--- a/.hgignore Wed Jan 04 14:26:08 2017 +0100
+++ b/.hgignore Wed Jan 04 14:26:19 2017 +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/.hgtags Wed Jan 04 14:26:08 2017 +0100
+++ b/.hgtags Wed Jan 04 14:26:19 2017 +0100
@@ -32,8 +32,4 @@
8f4a332500e41bb67efc3e141608829473606a72 Isabelle2014
5ae2a2e74c93eafeb00b1ddeef0404256745ebba Isabelle2015
d3996d5873ddcf1115ec8d3d511a0bb5dbd1cfc4 Isabelle2016
-666c7475f4f7e9ba46c59170026230787c504ca7 Isabelle2016-1-RC0
-9ee2480d10b7404683aa7f4c3a30d44cbb6a21b9 Isabelle2016-1-RC1
-2bf4fdcebd495516947e5e85f3b3db01d5fbe1a4 Isabelle2016-1-RC2
-51be997d0698583bf2d3f5a99f37381a146d3a6c Isabelle2016-1-RC3
-49708cffb98dc6ced89f66b10662e6af2808bebd Isabelle2016-1-RC4
+7aa3c52f27aade2cada22206cf0477b30a25f781 Isabelle2016-1
--- a/ANNOUNCE Wed Jan 04 14:26:08 2017 +0100
+++ b/ANNOUNCE Wed Jan 04 14:26:19 2017 +0100
@@ -4,7 +4,7 @@
Isabelle2016-1 is now available.
This version introduces significant changes over Isabelle2016: see the NEWS
-file for further details. Some notable changes:
+file for further details. Some notable points:
* Improved Isabelle/jEdit Prover IDE: more support for formal text structure,
more visual feedback.
--- a/Admin/components/components.sha1 Wed Jan 04 14:26:08 2017 +0100
+++ b/Admin/components/components.sha1 Wed Jan 04 14:26:19 2017 +0100
@@ -138,6 +138,7 @@
5b70c12c95a90d858f90c1945011289944ea8e17 polyml-5.6-20160118.tar.gz
5b19dc93082803b82aa553a5cfb3e914606c0ffd polyml-5.6.tar.gz
853ab0e9ff2b73790cc80a2d36cbff8b03e50a8e polyml-test-7a7b742897e9.tar.gz
+c629cd499a724bbe37b962f727e4ff340c50299d polyml-test-8529546198aa.tar.gz
8ee375cfc38972f080dbc78f07b68dac03efe968 ProofGeneral-3.7.1.1.tar.gz
847b52c0676b5eb0fbf0476f64fc08c2d72afd0c ProofGeneral-4.1.tar.gz
8e0b2b432755ef11d964e20637d1bc567d1c0477 ProofGeneral-4.2-1.tar.gz
--- a/Admin/polyml/README Wed Jan 04 14:26:08 2017 +0100
+++ b/Admin/polyml/README Wed Jan 04 14:26:19 2017 +0100
@@ -1,32 +1,27 @@
Poly/ML for Isabelle
====================
-This compilation of Poly/ML 5.6 (http://www.polyml.org) is based on the source
-distribution from https://github.com/polyml/polyml/releases/tag/v5.6/.
+This compilation of Poly/ML (http://www.polyml.org) is based on the repository
+version https://github.com/polyml/polyml/commit/8529546198aa
-On Linux the sources have changed as follows, in order to evade a
-potential conflict of /bin/bash versus /bin/sh -> dash (notably on
-Ubuntu and Debian):
+The Isabelle repository provides the administrative tool "build_polyml", which
+can be used in the polyml component directory as follows.
-diff -r src-orig/libpolyml/process_env.cpp src/libpolyml/process_env.cpp
-228c228
-< execve("/bin/sh", argv, environ);
----
-> execvp("bash", argv);
+* Linux:
+ isabelle build_polyml -m32 -s sha1 src --with-gmp
+ isabelle build_polyml -m64 -s sha1 src --with-gmp
-The included build script is used like this:
+* Mac OS X:
- ./build src x86-linux --with-gmp
- ./build src x86_64-linux --with-gmp
- ./build src x86-darwin --without-gmp
- ./build src x86_64-darwin --without-gmp
- ./build src x86-windows --with-gmp
- ./build src x86_64-windows --with-gmp
+ isabelle build_polyml -m32 -s sha1 src --without-gmp
+ isabelle build_polyml -m64 -s sha1 src --without-gmp
-Also note that the separate "sha1" library module is required for
-efficient digestion of strings according to SHA-1.
+* Windows (Cygwin shell)
+
+ isabelle build_polyml -M /cygdrive/c/msys64 -m32 -s sha1 src --with-gmp
+ isabelle build_polyml -M /cygdrive/c/msys64 -m64 -s sha1 src --with-gmp
Makarius
- 11-Feb-2016
+ 10-Dec-2016
--- a/Admin/polyml/build Wed Jan 04 14:26:08 2017 +0100
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,117 +0,0 @@
-#!/usr/bin/env bash
-#
-# Multi-platform build script for Poly/ML
-
-THIS="$(cd "$(dirname "$0")"; pwd)"
-PRG="$(basename "$0")"
-
-
-# diagnostics
-
-function usage()
-{
- echo
- echo "Usage: $PRG SOURCE TARGET [OPTIONS]"
- echo
- echo " Build Poly/ML in SOURCE directory for given platform in TARGET,"
- echo " using the usual Isabelle platform identifiers."
- echo
- echo " Additional options for ./configure may be given, e.g. --with-gmp"
- echo
- exit 1
-}
-
-function fail()
-{
- echo "$1" >&2
- exit 2
-}
-
-
-# command line args
-
-[ "$#" -eq 0 ] && usage
-SOURCE="$1"; shift
-
-[ "$#" -eq 0 ] && usage
-TARGET="$1"; shift
-
-USER_OPTIONS=("$@")
-
-
-# main
-
-[ -d "$SOURCE" ] || fail "Bad source directory: \"$SOURCE\""
-
-case "$TARGET" in
- x86-linux)
- OPTIONS=()
- ;;
- x86_64-linux)
- OPTIONS=()
- ;;
- x86-darwin)
- OPTIONS=(--build=i686-darwin CFLAGS='-arch i686 -O3 -I../libffi/include'
- CXXFLAGS='-arch i686 -O3 -I../libffi/include' CCASFLAGS='-arch i686 -O3'
- LDFLAGS='-segprot POLY rwx rwx')
- ;;
- x86_64-darwin)
- OPTIONS=(--build=x86_64-darwin CFLAGS='-arch x86_64 -O3 -I../libffi/include'
- CXXFLAGS='-arch x86_64 -O3 -I../libffi/include' CCASFLAGS='-arch x86_64'
- LDFLAGS='-segprot POLY rwx rwx')
- ;;
- x86-cygwin)
- OPTIONS=()
- ;;
- x86-windows)
- OPTIONS=(--host=i686-w32-mingw32 CPPFLAGS='-I/mingw32/include' --disable-windows-gui)
- PATH="/mingw32/bin:$PATH"
- ;;
- x86_64-windows)
- OPTIONS=(--host=x86_64-w64-mingw32 CPPFLAGS='-I/mingw64/include' --disable-windows-gui)
- PATH="/mingw64/bin:$PATH"
- ;;
- *)
- fail "Bad platform identifier: \"$TARGET\""
- ;;
-esac
-
-(
- cd "$SOURCE"
- make distclean
-
- { ./configure --prefix="$PWD/$TARGET" "${OPTIONS[@]}" --enable-intinf-as-int "${USER_OPTIONS[@]}" && \
- make compiler && \
- make compiler && \
- make install; } || fail "Build failed"
-)
-
-mkdir -p "$TARGET"
-for X in "$TARGET"/*
-do
- [ -d "$X" ] && rm -rf "$X"
-done
-rm -rf "$TARGET/polyml"
-cp -a "$THIS/polyi" "$TARGET/"
-mv "$SOURCE/$TARGET/bin/"* "$TARGET/"
-mv "$SOURCE/$TARGET/lib/"* "$TARGET/"
-rmdir "$SOURCE/$TARGET/bin" "$SOURCE/$TARGET/lib"
-rm -rf "$SOURCE/$TARGET/share"
-
-case "$TARGET" in
- x86-cygwin)
- peflags -x8192000 -z500 "$TARGET/poly.exe"
- ;;
- x86-windows)
- for X in libgcc_s_dw2-1.dll libgmp-10.dll libstdc++-6.dll
- do
- cp "/mingw32/bin/$X" "$TARGET/."
- done
- ;;
- x86_64-windows)
- for X in libgcc_s_seh-1.dll libgmp-10.dll libstdc++-6.dll
- do
- cp "/mingw64/bin/$X" "$TARGET/."
- done
- ;;
-esac
--- a/Admin/polyml/settings Wed Jan 04 14:26:08 2017 +0100
+++ b/Admin/polyml/settings Wed Jan 04 14:26:19 2017 +0100
@@ -36,12 +36,12 @@
do
if [ -z "$ML_HOME" ]
then
- if "$POLYML_HOME/$PLATFORM/polyml" -v </dev/null >/dev/null 2>/dev/null
+ if "$POLYML_HOME/$PLATFORM/polyi" -v </dev/null >/dev/null 2>/dev/null
then
# ML settings
- ML_SYSTEM=polyml-5.6
+ ML_SYSTEM=polyml-5.7
ML_PLATFORM="$PLATFORM"
ML_HOME="$POLYML_HOME/$ML_PLATFORM"
ML_SOURCES="$POLYML_HOME/src"
--- a/CONTRIBUTORS Wed Jan 04 14:26:08 2017 +0100
+++ b/CONTRIBUTORS Wed Jan 04 14:26:19 2017 +0100
@@ -10,6 +10,10 @@
Contributions to Isabelle2016-1
-------------------------------
+* December 2016: Ondřej Kunčar, TUM
+ Types_To_Sets: experimental extension of Higher-Order Logic to allow
+ translation of types to sets.
+
* October 2016: Jasmin Blanchette
Integration of Nunchaku model finder.
--- a/NEWS Wed Jan 04 14:26:08 2017 +0100
+++ b/NEWS Wed Jan 04 14:26:19 2017 +0100
@@ -3,14 +3,49 @@
(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 ***
+
+* Dropped aliasses RangeP, DomainP for Rangep, Domainp respectively.
+INCOMPATIBILITY.
+
+* Dropped abbreviation transP, antisymP, single_valuedP;
+use constants transp, antisymp, single_valuedp instead.
+INCOMPATIBILITY.
+
+* 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
it no longer fails on nested datatypes. Slight INCOMPATIBILITY.
+* The theorem in Permutations has been renamed:
+ bij_swap_ompose_bij ~> bij_swap_compose_bij
New in Isabelle2016-1 (December 2016)
@@ -947,6 +982,9 @@
* Session Old_Number_Theory has been removed, after porting remaining
theories.
+* Session HOL-Types_To_Sets provides an experimental extension of
+Higher-Order Logic to allow translation of types to sets.
+
*** ML ***
--- a/etc/components Wed Jan 04 14:26:08 2017 +0100
+++ b/etc/components Wed Jan 04 14:26:19 2017 +0100
@@ -1,6 +1,7 @@
#hard-wired components
src/Tools/jEdit
src/Tools/Graphview
+src/Tools/VSCode
src/HOL/Mirabelle
src/HOL/Mutabelle
src/HOL/Library/Sum_of_Squares
--- a/etc/symbols Wed Jan 04 14:26:08 2017 +0100
+++ b/etc/symbols Wed Jan 04 14:26:19 2017 +0100
@@ -361,7 +361,7 @@
\<comment> code: 0x002015 group: document font: IsabelleText
\<open> code: 0x002039 group: punctuation font: IsabelleText abbrev: <<
\<close> code: 0x00203a group: punctuation font: IsabelleText abbrev: >>
-\<here> code: 0x002302 font: IsabelleText
+\<^here> code: 0x002302 font: IsabelleText
\<^undefined> code: 0x002756 font: IsabelleText
\<^noindent> code: 0x0021e4 group: document font: IsabelleText
\<^smallskip> code: 0x002508 group: document font: IsabelleText
--- a/lib/texinputs/isabellesym.sty Wed Jan 04 14:26:08 2017 +0100
+++ b/lib/texinputs/isabellesym.sty Wed Jan 04 14:26:19 2017 +0100
@@ -370,3 +370,4 @@
\newcommand{\isactrlundefined}{\isakeyword{undefined}\ }
\newcommand{\isactrlfile}{\isakeyword{file}\ }
\newcommand{\isactrldir}{\isakeyword{dir}\ }
+\newcommand{\isactrlhere}{\isakeyword{here}\ }
--- a/src/Doc/JEdit/JEdit.thy Wed Jan 04 14:26:08 2017 +0100
+++ b/src/Doc/JEdit/JEdit.thy Wed Jan 04 14:26:19 2017 +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
@@ -723,6 +729,11 @@
The above options are accessible in the menu \<^emph>\<open>Plugins / Plugin Options /
Isabelle / General\<close>.
+
+ \<^medskip> Automatic indentation has a tendency to produce trailing whitespace. That
+ can be purged manually with the jEdit action @{action "remove-trailing-ws"}
+ (shortcut \<^verbatim>\<open>C+e r\<close>). Moreover, the \<^emph>\<open>WhiteSpace\<close> plugin provides some
+ aggressive options to trim whitespace on buffer-save.
\<close>
--- a/src/HOL/Algebra/Coset.thy Wed Jan 04 14:26:08 2017 +0100
+++ b/src/HOL/Algebra/Coset.thy Wed Jan 04 14:26:19 2017 +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 Wed Jan 04 14:26:08 2017 +0100
+++ b/src/HOL/Algebra/Divisibility.thy Wed Jan 04 14:26:19 2017 +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/Cauchy_Integral_Theorem.thy Wed Jan 04 14:26:08 2017 +0100
+++ b/src/HOL/Analysis/Cauchy_Integral_Theorem.thy Wed Jan 04 14:26:19 2017 +0100
@@ -4601,11 +4601,11 @@
by metis
note ee_rule = ee [THEN conjunct2, rule_format]
define C where "C = (\<lambda>t. ball t (ee t / 3)) ` {0..1}"
- have "\<forall>t \<in> C. open t" by (simp add: C_def)
- moreover have "{0..1} \<subseteq> \<Union>C"
- using ee [THEN conjunct1] by (auto simp: C_def dist_norm)
- ultimately obtain C' where C': "C' \<subseteq> C" "finite C'" and C'01: "{0..1} \<subseteq> \<Union>C'"
- by (rule compactE [OF compact_interval])
+ obtain C' where C': "C' \<subseteq> C" "finite C'" and C'01: "{0..1} \<subseteq> \<Union>C'"
+ proof (rule compactE [OF compact_interval])
+ show "{0..1} \<subseteq> \<Union>C"
+ using ee [THEN conjunct1] by (auto simp: C_def dist_norm)
+ qed (use C_def in auto)
define kk where "kk = {t \<in> {0..1}. ball t (ee t / 3) \<in> C'}"
have kk01: "kk \<subseteq> {0..1}" by (auto simp: kk_def)
define e where "e = Min (ee ` kk)"
--- a/src/HOL/Analysis/Complex_Transcendental.thy Wed Jan 04 14:26:08 2017 +0100
+++ b/src/HOL/Analysis/Complex_Transcendental.thy Wed Jan 04 14:26:19 2017 +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/Analysis/Ordered_Euclidean_Space.thy Wed Jan 04 14:26:08 2017 +0100
+++ b/src/HOL/Analysis/Ordered_Euclidean_Space.thy Wed Jan 04 14:26:19 2017 +0100
@@ -197,13 +197,13 @@
shows "closed {a..}"
by (simp add: eucl_le_atLeast[symmetric])
-lemma bounded_closed_interval:
+lemma bounded_closed_interval [simp]:
fixes a :: "'a::ordered_euclidean_space"
shows "bounded {a .. b}"
using bounded_cbox[of a b]
by (metis interval_cbox)
-lemma convex_closed_interval:
+lemma convex_closed_interval [simp]:
fixes a :: "'a::ordered_euclidean_space"
shows "convex {a .. b}"
using convex_box[of a b]
@@ -214,11 +214,11 @@
using image_smult_cbox[of m a b]
by (simp add: cbox_interval)
-lemma is_interval_closed_interval:
+lemma is_interval_closed_interval [simp]:
"is_interval {a .. (b::'a::ordered_euclidean_space)}"
by (metis cbox_interval is_interval_cbox)
-lemma compact_interval:
+lemma compact_interval [simp]:
fixes a b::"'a::ordered_euclidean_space"
shows "compact {a .. b}"
by (metis compact_cbox interval_cbox)
--- a/src/HOL/Analysis/Topology_Euclidean_Space.thy Wed Jan 04 14:26:08 2017 +0100
+++ b/src/HOL/Analysis/Topology_Euclidean_Space.thy Wed Jan 04 14:26:19 2017 +0100
@@ -794,28 +794,28 @@
by (simp add: subtopology_superset)
lemma openin_subtopology_empty:
- "openin (subtopology U {}) s \<longleftrightarrow> s = {}"
+ "openin (subtopology U {}) S \<longleftrightarrow> S = {}"
by (metis Int_empty_right openin_empty openin_subtopology)
lemma closedin_subtopology_empty:
- "closedin (subtopology U {}) s \<longleftrightarrow> s = {}"
+ "closedin (subtopology U {}) S \<longleftrightarrow> S = {}"
by (metis Int_empty_right closedin_empty closedin_subtopology)
-lemma closedin_subtopology_refl:
- "closedin (subtopology U u) u \<longleftrightarrow> u \<subseteq> topspace U"
+lemma closedin_subtopology_refl [simp]:
+ "closedin (subtopology U X) X \<longleftrightarrow> X \<subseteq> topspace U"
by (metis closedin_def closedin_topspace inf.absorb_iff2 le_inf_iff topspace_subtopology)
lemma openin_imp_subset:
- "openin (subtopology U s) t \<Longrightarrow> t \<subseteq> s"
+ "openin (subtopology U S) T \<Longrightarrow> T \<subseteq> S"
by (metis Int_iff openin_subtopology subsetI)
lemma closedin_imp_subset:
- "closedin (subtopology U s) t \<Longrightarrow> t \<subseteq> s"
+ "closedin (subtopology U S) T \<Longrightarrow> T \<subseteq> S"
by (simp add: closedin_def topspace_subtopology)
lemma openin_subtopology_Un:
- "openin (subtopology U t) s \<and> openin (subtopology U u) s
- \<Longrightarrow> openin (subtopology U (t \<union> u)) s"
+ "openin (subtopology U T) S \<and> openin (subtopology U u) S
+ \<Longrightarrow> openin (subtopology U (T \<union> u)) S"
by (simp add: openin_subtopology) blast
@@ -1061,6 +1061,9 @@
lemma sphere_cball [simp,intro]: "sphere z r \<subseteq> cball z r"
by force
+lemma cball_diff_sphere: "cball a r - sphere a r = ball a r"
+ by auto
+
lemma subset_ball[intro]: "d \<le> e \<Longrightarrow> ball x d \<subseteq> ball x e"
by (simp add: subset_eq)
@@ -1073,6 +1076,12 @@
lemma ball_min_Int: "ball a (min r s) = ball a r \<inter> ball a s"
by (simp add: set_eq_iff)
+lemma cball_max_Un: "cball a (max r s) = cball a r \<union> cball a s"
+ by (simp add: set_eq_iff) arith
+
+lemma cball_min_Int: "cball a (min r s) = cball a r \<inter> cball a s"
+ by (simp add: set_eq_iff)
+
lemma cball_diff_eq_sphere: "cball a r - ball a r = {x. dist x a = r}"
by (auto simp: cball_def ball_def dist_commute)
@@ -2463,6 +2472,54 @@
apply (simp add: connected_component_maximal connected_component_mono subset_antisym)
using connected_component_eq_empty by blast
+proposition connected_Times:
+ assumes S: "connected S" and T: "connected T"
+ shows "connected (S \<times> T)"
+proof (clarsimp simp add: connected_iff_connected_component)
+ fix x y x' y'
+ assume xy: "x \<in> S" "y \<in> T" "x' \<in> S" "y' \<in> T"
+ with xy obtain U V where U: "connected U" "U \<subseteq> S" "x \<in> U" "x' \<in> U"
+ and V: "connected V" "V \<subseteq> T" "y \<in> V" "y' \<in> V"
+ using S T \<open>x \<in> S\<close> \<open>x' \<in> S\<close> by blast+
+ show "connected_component (S \<times> T) (x, y) (x', y')"
+ unfolding connected_component_def
+ proof (intro exI conjI)
+ show "connected ((\<lambda>x. (x, y)) ` U \<union> Pair x' ` V)"
+ proof (rule connected_Un)
+ have "continuous_on U (\<lambda>x. (x, y))"
+ by (intro continuous_intros)
+ then show "connected ((\<lambda>x. (x, y)) ` U)"
+ by (rule connected_continuous_image) (rule \<open>connected U\<close>)
+ have "continuous_on V (Pair x')"
+ by (intro continuous_intros)
+ then show "connected (Pair x' ` V)"
+ by (rule connected_continuous_image) (rule \<open>connected V\<close>)
+ qed (use U V in auto)
+ qed (use U V in auto)
+qed
+
+corollary connected_Times_eq [simp]:
+ "connected (S \<times> T) \<longleftrightarrow> S = {} \<or> T = {} \<or> connected S \<and> connected T" (is "?lhs = ?rhs")
+proof
+ assume L: ?lhs
+ show ?rhs
+ proof (cases "S = {} \<or> T = {}")
+ case True
+ then show ?thesis by auto
+ next
+ case False
+ have "connected (fst ` (S \<times> T))" "connected (snd ` (S \<times> T))"
+ using continuous_on_fst continuous_on_snd continuous_on_id
+ by (blast intro: connected_continuous_image [OF _ L])+
+ with False show ?thesis
+ by auto
+ qed
+next
+ assume ?rhs
+ then show ?lhs
+ by (auto simp: connected_Times)
+qed
+
subsection \<open>The set of connected components of a set\<close>
@@ -7240,7 +7297,7 @@
then have "\<forall>c\<in>{T. open T \<and> S \<inter> T \<in> C}. open c" and "S \<subseteq> \<Union>{T. open T \<and> S \<inter> T \<in> C}"
unfolding openin_open by force+
with \<open>compact S\<close> obtain D where "D \<subseteq> {T. open T \<and> S \<inter> T \<in> C}" and "finite D" and "S \<subseteq> \<Union>D"
- by (rule compactE)
+ by (meson compactE)
then have "image (\<lambda>T. S \<inter> T) D \<subseteq> C \<and> finite (image (\<lambda>T. S \<inter> T) D) \<and> S \<subseteq> \<Union>(image (\<lambda>T. S \<inter> T) D)"
by auto
then show "\<exists>D\<subseteq>C. finite D \<and> S \<subseteq> \<Union>D" ..
@@ -10189,7 +10246,7 @@
by metis
from * have "\<forall>t\<in>Y ` K. open t" "K \<subseteq> \<Union>(Y ` K)" by auto
with \<open>compact K\<close> obtain CC where CC: "CC \<subseteq> Y ` K" "finite CC" "K \<subseteq> \<Union>CC"
- by (rule compactE)
+ by (meson compactE)
then obtain c where c: "\<And>C. C \<in> CC \<Longrightarrow> c C \<in> K \<and> C = Y (c C)"
by (force intro!: choice)
with * CC show ?thesis
--- a/src/HOL/Code_Numeral.thy Wed Jan 04 14:26:08 2017 +0100
+++ b/src/HOL/Code_Numeral.thy Wed Jan 04 14:26:19 2017 +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/approximation.ML Wed Jan 04 14:26:08 2017 +0100
+++ b/src/HOL/Decision_Procs/approximation.ML Wed Jan 04 14:26:19 2017 +0100
@@ -43,7 +43,7 @@
end
fun approximation_conv ctxt ct =
- approximation_oracle (Proof_Context.theory_of ctxt, Thm.term_of ct |> tap (tracing o Syntax.string_of_term ctxt));
+ approximation_oracle (Proof_Context.theory_of ctxt, Thm.term_of ct);
fun approximate ctxt t =
approximation_oracle (Proof_Context.theory_of ctxt, t)
--- a/src/HOL/Decision_Procs/cooper_tac.ML Wed Jan 04 14:26:08 2017 +0100
+++ b/src/HOL/Decision_Procs/cooper_tac.ML Wed Jan 04 14:26:19 2017 +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 Wed Jan 04 14:26:08 2017 +0100
+++ b/src/HOL/Decision_Procs/mir_tac.ML Wed Jan 04 14:26:19 2017 +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 Wed Jan 04 14:26:08 2017 +0100
+++ b/src/HOL/Divides.thy Wed Jan 04 14:26:19 2017 +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
@@ -783,7 +781,38 @@
lemma one_mod_numeral [simp]:
"1 mod numeral n = snd (divmod num.One n)"
by (simp add: snd_divmod)
-
+
+text \<open>Computing congruences modulo \<open>2 ^ q\<close>\<close>
+
+lemma cong_exp_iff_simps:
+ "numeral n mod numeral Num.One = 0
+ \<longleftrightarrow> True"
+ "numeral (Num.Bit0 n) mod numeral (Num.Bit0 q) = 0
+ \<longleftrightarrow> numeral n mod numeral q = 0"
+ "numeral (Num.Bit1 n) mod numeral (Num.Bit0 q) = 0
+ \<longleftrightarrow> False"
+ "numeral m mod numeral Num.One = (numeral n mod numeral Num.One)
+ \<longleftrightarrow> True"
+ "numeral Num.One mod numeral (Num.Bit0 q) = (numeral Num.One mod numeral (Num.Bit0 q))
+ \<longleftrightarrow> True"
+ "numeral Num.One mod numeral (Num.Bit0 q) = (numeral (Num.Bit0 n) mod numeral (Num.Bit0 q))
+ \<longleftrightarrow> False"
+ "numeral Num.One mod numeral (Num.Bit0 q) = (numeral (Num.Bit1 n) mod numeral (Num.Bit0 q))
+ \<longleftrightarrow> (numeral n mod numeral q) = 0"
+ "numeral (Num.Bit0 m) mod numeral (Num.Bit0 q) = (numeral Num.One mod numeral (Num.Bit0 q))
+ \<longleftrightarrow> False"
+ "numeral (Num.Bit0 m) mod numeral (Num.Bit0 q) = (numeral (Num.Bit0 n) mod numeral (Num.Bit0 q))
+ \<longleftrightarrow> numeral m mod numeral q = (numeral n mod numeral q)"
+ "numeral (Num.Bit0 m) mod numeral (Num.Bit0 q) = (numeral (Num.Bit1 n) mod numeral (Num.Bit0 q))
+ \<longleftrightarrow> False"
+ "numeral (Num.Bit1 m) mod numeral (Num.Bit0 q) = (numeral Num.One mod numeral (Num.Bit0 q))
+ \<longleftrightarrow> (numeral m mod numeral q) = 0"
+ "numeral (Num.Bit1 m) mod numeral (Num.Bit0 q) = (numeral (Num.Bit0 n) mod numeral (Num.Bit0 q))
+ \<longleftrightarrow> False"
+ "numeral (Num.Bit1 m) mod numeral (Num.Bit0 q) = (numeral (Num.Bit1 n) mod numeral (Num.Bit0 q))
+ \<longleftrightarrow> numeral m mod numeral q = (numeral n mod numeral q)"
+ by (auto simp add: case_prod_beta dest: arg_cong [of _ _ even])
+
end
@@ -799,18 +828,18 @@
@{term "q::nat"}(uotient) and @{term "r::nat"}(emainder).
\<close>
-definition divmod_nat_rel :: "nat \<Rightarrow> nat \<Rightarrow> nat \<times> nat \<Rightarrow> bool" where
- "divmod_nat_rel m n qr \<longleftrightarrow>
- m = fst qr * n + snd qr \<and>
- (if n = 0 then fst qr = 0 else if n > 0 then 0 \<le> snd qr \<and> snd qr < n else n < snd qr \<and> snd qr \<le> 0)"
-
-text \<open>@{const divmod_nat_rel} is total:\<close>
-
-qualified lemma divmod_nat_rel_ex:
- obtains q r where "divmod_nat_rel m n (q, r)"
+inductive eucl_rel_nat :: "nat \<Rightarrow> nat \<Rightarrow> nat \<times> nat \<Rightarrow> bool"
+ where eucl_rel_nat_by0: "eucl_rel_nat m 0 (0, m)"
+ | eucl_rel_natI: "r < n \<Longrightarrow> m = q * n + r \<Longrightarrow> eucl_rel_nat m n (q, r)"
+
+text \<open>@{const eucl_rel_nat} is total:\<close>
+
+qualified lemma eucl_rel_nat_ex:
+ obtains q r where "eucl_rel_nat m n (q, r)"
proof (cases "n = 0")
- case True with that show thesis
- by (auto simp add: divmod_nat_rel_def)
+ case True
+ with that eucl_rel_nat_by0 show thesis
+ by blast
next
case False
have "\<exists>q r. m = q * n + r \<and> r < n"
@@ -826,92 +855,124 @@
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
qed
qed
- with that show thesis
- using \<open>n \<noteq> 0\<close> by (auto simp add: divmod_nat_rel_def)
+ with that \<open>n \<noteq> 0\<close> eucl_rel_natI show thesis
+ by blast
qed
-text \<open>@{const divmod_nat_rel} is injective:\<close>
-
-qualified lemma divmod_nat_rel_unique:
- assumes "divmod_nat_rel m n qr"
- and "divmod_nat_rel m n qr'"
- shows "qr = qr'"
+text \<open>@{const eucl_rel_nat} is injective:\<close>
+
+qualified lemma eucl_rel_nat_unique_div:
+ assumes "eucl_rel_nat m n (q, r)"
+ and "eucl_rel_nat m n (q', r')"
+ shows "q = q'"
proof (cases "n = 0")
case True with assms show ?thesis
- by (cases qr, cases qr')
- (simp add: divmod_nat_rel_def)
+ by (auto elim: eucl_rel_nat.cases)
next
case False
- have aux: "\<And>q r q' r'. q' * n + r' = q * n + r \<Longrightarrow> r < n \<Longrightarrow> q' \<le> (q::nat)"
- apply (rule leI)
- apply (subst less_iff_Suc_add)
- 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)
- with assms have "snd qr = snd qr'"
- by (simp add: divmod_nat_rel_def)
- with * show ?thesis by (cases qr, cases qr') simp
+ have *: "q' \<le> q" if "q' * n + r' = q * n + r" "r < n" for q r q' r' :: nat
+ proof (rule ccontr)
+ assume "\<not> q' \<le> q"
+ then have "q < q'"
+ by (simp add: not_le)
+ with that show False
+ by (auto simp add: less_iff_Suc_add algebra_simps)
+ qed
+ from \<open>n \<noteq> 0\<close> assms show ?thesis
+ by (auto intro: order_antisym elim: eucl_rel_nat.cases dest: * sym split: if_splits)
+qed
+
+qualified lemma eucl_rel_nat_unique_mod:
+ assumes "eucl_rel_nat m n (q, r)"
+ and "eucl_rel_nat m n (q', r')"
+ shows "r = r'"
+proof -
+ from assms have "q' = q"
+ by (auto intro: eucl_rel_nat_unique_div)
+ with assms show ?thesis
+ by (auto elim!: eucl_rel_nat.cases)
qed
text \<open>
We instantiate divisibility on the natural numbers by
- means of @{const divmod_nat_rel}:
+ means of @{const eucl_rel_nat}:
\<close>
qualified definition divmod_nat :: "nat \<Rightarrow> nat \<Rightarrow> nat \<times> nat" where
- "divmod_nat m n = (THE qr. divmod_nat_rel m n qr)"
-
-qualified lemma divmod_nat_rel_divmod_nat:
- "divmod_nat_rel m n (divmod_nat m n)"
+ "divmod_nat m n = (THE qr. eucl_rel_nat m n qr)"
+
+qualified lemma eucl_rel_nat_divmod_nat:
+ "eucl_rel_nat m n (divmod_nat m n)"
proof -
- from divmod_nat_rel_ex
- obtain qr where rel: "divmod_nat_rel m n qr" .
+ from eucl_rel_nat_ex
+ obtain q r where rel: "eucl_rel_nat m n (q, r)" .
then show ?thesis
- by (auto simp add: divmod_nat_def intro: theI elim: divmod_nat_rel_unique)
+ by (auto simp add: divmod_nat_def intro: theI
+ elim: eucl_rel_nat_unique_div eucl_rel_nat_unique_mod)
qed
qualified lemma divmod_nat_unique:
- assumes "divmod_nat_rel m n qr"
- shows "divmod_nat m n = qr"
- 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)
-
-qualified lemma divmod_nat_zero_left: "divmod_nat 0 n = (0, 0)"
- by (simp add: Divides.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)
+ "divmod_nat m n = (q, r)" if "eucl_rel_nat m n (q, r)"
+ using that
+ by (auto simp add: divmod_nat_def intro: eucl_rel_nat_divmod_nat elim: eucl_rel_nat_unique_div eucl_rel_nat_unique_mod)
+
+qualified lemma divmod_nat_zero:
+ "divmod_nat m 0 = (0, m)"
+ by (rule divmod_nat_unique) (fact eucl_rel_nat_by0)
+
+qualified lemma divmod_nat_zero_left:
+ "divmod_nat 0 n = (0, 0)"
+ by (rule divmod_nat_unique)
+ (cases n, auto intro: eucl_rel_nat_by0 eucl_rel_natI)
+
+qualified lemma divmod_nat_base:
+ "m < n \<Longrightarrow> divmod_nat m n = (0, m)"
+ by (rule divmod_nat_unique)
+ (cases n, auto intro: eucl_rel_nat_by0 eucl_rel_natI)
qualified lemma divmod_nat_step:
assumes "0 < n" and "n \<le> m"
- shows "divmod_nat m n = apfst Suc (divmod_nat (m - n) n)"
+ shows "divmod_nat m n =
+ (Suc (fst (divmod_nat (m - n) n)), snd (divmod_nat (m - n) n))"
proof (rule divmod_nat_unique)
- 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
+ have "eucl_rel_nat (m - n) n (divmod_nat (m - n) n)"
+ by (fact eucl_rel_nat_divmod_nat)
+ then show "eucl_rel_nat m n (Suc
+ (fst (divmod_nat (m - n) n)), snd (divmod_nat (m - n) n))"
+ using assms
+ by (auto split: if_splits intro: eucl_rel_natI elim!: eucl_rel_nat.cases 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"
@@ -926,17 +987,20 @@
by (simp add: prod_eq_iff)
lemma div_nat_unique:
- assumes "divmod_nat_rel m n (q, r)"
+ assumes "eucl_rel_nat 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)"
+ assumes "eucl_rel_nat m n (q, r)"
shows "m mod n = r"
- 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 assms
+ by (auto dest!: Divides.divmod_nat_unique simp add: prod_eq_iff)
+
+lemma eucl_rel_nat: "eucl_rel_nat m n (m div n, m mod n)"
+ using Divides.eucl_rel_nat_divmod_nat
+ by (simp add: divmod_nat_div_mod)
text \<open>The ''recursion'' equations for @{const divide} and @{const modulo}\<close>
@@ -964,11 +1028,39 @@
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 eucl_rel_nat [of m n]
+ by (auto elim: eucl_rel_nat.cases)
+
+lemma mod_le_divisor [simp]:
+ fixes m n :: nat
+ assumes "n > 0"
+ shows "m mod n \<le> n"
+ using assms eucl_rel_nat [of m n]
+ by (auto elim: eucl_rel_nat.cases)
+
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)
+ using eucl_rel_nat [of m n]
+ by (auto elim: eucl_rel_nat.cases)
+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 intro!: eucl_rel_natI 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 +1068,42 @@
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)
+ show "(m * n) div (m * q) = n div q"
+ proof (cases "q = 0")
+ case True
+ then show ?thesis
+ by simp
+ next
+ case False
+ show ?thesis
+ proof (rule div_nat_unique [of _ _ _ "m * (n mod q)"])
+ show "eucl_rel_nat (m * n) (m * q) (n div q, m * (n mod q))"
+ by (rule eucl_rel_natI)
+ (use \<open>m \<noteq> 0\<close> \<open>q \<noteq> 0\<close> div_mult_mod_eq [of n q] in \<open>auto simp add: algebra_simps distrib_left [symmetric]\<close>)
+ qed
+ qed
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 +1134,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 +1187,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,41 +1210,36 @@
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>
-lemma divmod_nat_rel_mult1_eq:
- "divmod_nat_rel b c (q, r)
- \<Longrightarrow> divmod_nat_rel (a * b) c (a * q + a * r div c, a * r mod c)"
-by (auto simp add: split_ifs divmod_nat_rel_def algebra_simps)
-
lemma div_mult1_eq:
"(a * b) div c = a * (b div c) + a * (b mod c) div (c::nat)"
-by (blast intro: divmod_nat_rel_mult1_eq [THEN div_nat_unique] divmod_nat_rel)
-
-lemma divmod_nat_rel_add1_eq:
- "divmod_nat_rel a c (aq, ar) \<Longrightarrow> divmod_nat_rel b c (bq, br)
- \<Longrightarrow> divmod_nat_rel (a + b) c (aq + bq + (ar + br) div c, (ar + br) mod c)"
-by (auto simp add: split_ifs divmod_nat_rel_def algebra_simps)
+ by (cases "c = 0")
+ (auto simp add: algebra_simps distrib_left [symmetric]
+ intro!: div_nat_unique [of _ _ _ "(a * (b mod c)) mod c"] eucl_rel_natI)
+
+lemma eucl_rel_nat_add1_eq:
+ "eucl_rel_nat a c (aq, ar) \<Longrightarrow> eucl_rel_nat b c (bq, br)
+ \<Longrightarrow> eucl_rel_nat (a + b) c (aq + bq + (ar + br) div c, (ar + br) mod c)"
+ by (auto simp add: split_ifs algebra_simps elim!: eucl_rel_nat.cases intro: eucl_rel_nat_by0 eucl_rel_natI)
(*NOT suitable for rewriting: the RHS has an instance of the LHS*)
lemma div_add1_eq:
- "(a+b) div (c::nat) = a div c + b div c + ((a mod c + b mod c) div c)"
-by (blast intro: divmod_nat_rel_add1_eq [THEN div_nat_unique] divmod_nat_rel)
-
-lemma divmod_nat_rel_mult2_eq:
- assumes "divmod_nat_rel a b (q, r)"
- shows "divmod_nat_rel a (b * c) (q div c, b *(q mod c) + r)"
-proof -
- { assume "r < b" and "0 < c"
- then have "b * (q mod c) + r < b * c"
+ "(a + b) div (c::nat) = a div c + b div c + ((a mod c + b mod c) div c)"
+by (blast intro: eucl_rel_nat_add1_eq [THEN div_nat_unique] eucl_rel_nat)
+
+lemma eucl_rel_nat_mult2_eq:
+ assumes "eucl_rel_nat a b (q, r)"
+ shows "eucl_rel_nat a (b * c) (q div c, b *(q mod c) + r)"
+proof (cases "c = 0")
+ case True
+ with assms show ?thesis
+ by (auto intro: eucl_rel_nat_by0 elim!: eucl_rel_nat.cases simp add: ac_simps)
+next
+ case False
+ { assume "r < b"
+ with False have "b * (q mod c) + r < b * c"
apply (cut_tac m = q and n = c in mod_less_divisor)
apply (drule_tac [2] m = "q mod c" in less_imp_Suc_add, auto)
apply (erule_tac P = "%x. lhs < rhs x" for lhs rhs in ssubst)
@@ -1147,15 +1247,15 @@
done
then have "r + b * (q mod c) < b * c"
by (simp add: ac_simps)
- } with assms show ?thesis
- by (auto simp add: divmod_nat_rel_def algebra_simps add_mult_distrib2 [symmetric])
+ } with assms False show ?thesis
+ by (auto simp add: algebra_simps add_mult_distrib2 [symmetric] elim!: eucl_rel_nat.cases intro: eucl_rel_nat.intros)
qed
lemma div_mult2_eq: "a div (b * c) = (a div b) div (c::nat)"
-by (force simp add: divmod_nat_rel [THEN divmod_nat_rel_mult2_eq, THEN div_nat_unique])
+by (force simp add: eucl_rel_nat [THEN eucl_rel_nat_mult2_eq, THEN div_nat_unique])
lemma mod_mult2_eq: "a mod (b * c) = b * (a div b mod c) + a mod (b::nat)"
-by (auto simp add: mult.commute divmod_nat_rel [THEN divmod_nat_rel_mult2_eq, THEN mod_nat_unique])
+by (auto simp add: mult.commute eucl_rel_nat [THEN eucl_rel_nat_mult2_eq, THEN mod_nat_unique])
instantiation nat :: semiring_numeral_div
begin
@@ -1180,25 +1280,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)"
@@ -1323,8 +1414,8 @@
from A B show ?lhs ..
next
assume P: ?lhs
- then have "divmod_nat_rel m n (q, m - n * q)"
- unfolding divmod_nat_rel_def by (auto simp add: ac_simps)
+ then have "eucl_rel_nat m n (q, m - n * q)"
+ by (auto intro: eucl_rel_natI simp add: ac_simps)
then have "m div n = q"
by (rule div_nat_unique)
then show ?rhs by simp
@@ -1519,11 +1610,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)
@@ -1609,43 +1695,25 @@
shows "Suc 0 mod numeral k = snd (divmod Num.One k)"
by (simp_all add: snd_divmod)
-lemma cut_eq_simps: \<comment> \<open>rewriting equivalence on \<open>n mod 2 ^ q\<close>\<close>
- fixes m n q :: num
- shows
- "numeral n mod numeral Num.One = (0::nat)
- \<longleftrightarrow> True"
- "numeral (Num.Bit0 n) mod numeral (Num.Bit0 q) = (0::nat)
- \<longleftrightarrow> numeral n mod numeral q = (0::nat)"
- "numeral (Num.Bit1 n) mod numeral (Num.Bit0 q) = (0::nat)
- \<longleftrightarrow> False"
- "numeral m mod numeral Num.One = (numeral n mod numeral Num.One :: nat)
- \<longleftrightarrow> True"
- "numeral Num.One mod numeral (Num.Bit0 q) = (numeral Num.One mod numeral (Num.Bit0 q) :: nat)
- \<longleftrightarrow> True"
- "numeral Num.One mod numeral (Num.Bit0 q) = (numeral (Num.Bit0 n) mod numeral (Num.Bit0 q) :: nat)
- \<longleftrightarrow> False"
- "numeral Num.One mod numeral (Num.Bit0 q) = (numeral (Num.Bit1 n) mod numeral (Num.Bit0 q) :: nat)
- \<longleftrightarrow> (numeral n mod numeral q :: nat) = 0"
- "numeral (Num.Bit0 m) mod numeral (Num.Bit0 q) = (numeral Num.One mod numeral (Num.Bit0 q) :: nat)
- \<longleftrightarrow> False"
- "numeral (Num.Bit0 m) mod numeral (Num.Bit0 q) = (numeral (Num.Bit0 n) mod numeral (Num.Bit0 q) :: nat)
- \<longleftrightarrow> numeral m mod numeral q = (numeral n mod numeral q :: nat)"
- "numeral (Num.Bit0 m) mod numeral (Num.Bit0 q) = (numeral (Num.Bit1 n) mod numeral (Num.Bit0 q) :: nat)
- \<longleftrightarrow> False"
- "numeral (Num.Bit1 m) mod numeral (Num.Bit0 q) = (numeral Num.One mod numeral (Num.Bit0 q) :: nat)
- \<longleftrightarrow> (numeral m mod numeral q :: nat) = 0"
- "numeral (Num.Bit1 m) mod numeral (Num.Bit0 q) = (numeral (Num.Bit0 n) mod numeral (Num.Bit0 q) :: nat)
- \<longleftrightarrow> False"
- "numeral (Num.Bit1 m) mod numeral (Num.Bit0 q) = (numeral (Num.Bit1 n) mod numeral (Num.Bit0 q) :: nat)
- \<longleftrightarrow> numeral m mod numeral q = (numeral n mod numeral q :: nat)"
- by (auto simp add: case_prod_beta Suc_double_not_eq_double double_not_eq_Suc_double)
-
subsection \<open>Division on @{typ int}\<close>
-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))"
+context
+begin
+
+inductive eucl_rel_int :: "int \<Rightarrow> int \<Rightarrow> int \<times> int \<Rightarrow> bool"
+ where eucl_rel_int_by0: "eucl_rel_int k 0 (0, k)"
+ | eucl_rel_int_dividesI: "l \<noteq> 0 \<Longrightarrow> k = q * l \<Longrightarrow> eucl_rel_int k l (q, 0)"
+ | eucl_rel_int_remainderI: "sgn r = sgn l \<Longrightarrow> \<bar>r\<bar> < \<bar>l\<bar>
+ \<Longrightarrow> k = q * l + r \<Longrightarrow> eucl_rel_int k l (q, r)"
+
+lemma eucl_rel_int_iff:
+ "eucl_rel_int k l (q, r) \<longleftrightarrow>
+ k = l * q + r \<and>
+ (if 0 < l then 0 \<le> r \<and> r < l else if l < 0 then l < r \<and> r \<le> 0 else q = 0)"
+ by (cases "r = 0")
+ (auto elim!: eucl_rel_int.cases intro: eucl_rel_int_by0 eucl_rel_int_dividesI eucl_rel_int_remainderI
+ simp add: ac_simps sgn_1_pos sgn_1_neg)
lemma unique_quotient_lemma:
"b * q' + r' \<le> b * q + r \<Longrightarrow> 0 \<le> r' \<Longrightarrow> r' < b \<Longrightarrow> r < b \<Longrightarrow> q' \<le> (q::int)"
@@ -1664,24 +1732,32 @@
by (rule_tac b = "-b" and r = "-r'" and r' = "-r" in unique_quotient_lemma) auto
lemma unique_quotient:
- "divmod_int_rel a b (q, r) \<Longrightarrow> divmod_int_rel a b (q', r') \<Longrightarrow> q = q'"
-apply (simp add: divmod_int_rel_def linorder_neq_iff split: if_split_asm)
-apply (blast intro: order_antisym
- dest: order_eq_refl [THEN unique_quotient_lemma]
- order_eq_refl [THEN unique_quotient_lemma_neg] sym)+
-done
+ "eucl_rel_int a b (q, r) \<Longrightarrow> eucl_rel_int a b (q', r') \<Longrightarrow> q = q'"
+ apply (simp add: eucl_rel_int_iff linorder_neq_iff split: if_split_asm)
+ apply (blast intro: order_antisym
+ dest: order_eq_refl [THEN unique_quotient_lemma]
+ order_eq_refl [THEN unique_quotient_lemma_neg] sym)+
+ done
lemma unique_remainder:
- "divmod_int_rel a b (q, r) \<Longrightarrow> divmod_int_rel a b (q', r') \<Longrightarrow> r = r'"
+ "eucl_rel_int a b (q, r) \<Longrightarrow> eucl_rel_int a b (q', r') \<Longrightarrow> r = r'"
apply (subgoal_tac "q = q'")
- apply (simp add: divmod_int_rel_def)
+ apply (simp add: eucl_rel_int_iff)
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,75 +1765,57 @@
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>
+lemma eucl_rel_int:
+ "eucl_rel_int k l (k div l, k mod l)"
+proof (cases k rule: int_cases3)
+ case zero
+ then show ?thesis
+ by (simp add: eucl_rel_int_iff 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
+ eucl_rel_int_iff 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
+ eucl_rel_int_iff divide_int_def modulo_int_def int_dvd_iff)
+qed
lemma divmod_int_unique:
- assumes "divmod_int_rel k l (q, r)"
+ assumes "eucl_rel_int k l (q, r)"
shows div_int_unique: "k div l = q" and mod_int_unique: "k mod l = r"
- using assms divmod_int_rel [of k l]
+ using assms eucl_rel_int [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 eucl_rel_int [of k l]
+ unfolding eucl_rel_int_iff 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)
+ fix k :: int show "k div 0 = 0"
+ by (rule div_int_unique, simp add: eucl_rel_int_iff)
next
- fix a :: int show "a 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: eucl_rel_int_iff ac_simps intro: div_int_unique [of _ _ _ 0])
+qed (simp_all add: sgn_mult mult_sgn_abs abs_sgn_eq)
end
@@ -1765,36 +1823,127 @@
"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
+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 div_abs_eq_div_nat:
+ "\<bar>k\<bar> div \<bar>l\<bar> = int (nat \<bar>k\<bar> div nat \<bar>l\<bar>)"
+ by (simp add: divide_int_def)
+
+lemma mod_abs_eq_div_nat:
+ "\<bar>k\<bar> mod \<bar>l\<bar> = int (nat \<bar>k\<bar> mod nat \<bar>l\<bar>)"
+ by (simp add: modulo_int_def dvd_int_unfold_dvd_nat)
+
+lemma div_sgn_abs_cancel:
+ fixes k l v :: int
+ assumes "v \<noteq> 0"
+ shows "(sgn v * \<bar>k\<bar>) div (sgn v * \<bar>l\<bar>) = \<bar>k\<bar> div \<bar>l\<bar>"
+proof -
+ from assms have "sgn v = - 1 \<or> sgn v = 1"
+ by (cases "v \<ge> 0") auto
+ then show ?thesis
+ using assms unfolding divide_int_def [of "sgn v * \<bar>k\<bar>" "sgn v * \<bar>l\<bar>"]
+ by (auto simp add: not_less div_abs_eq_div_nat)
+qed
+
+lemma div_eq_sgn_abs:
+ fixes k l v :: int
+ assumes "sgn k = sgn l"
+ shows "k div l = \<bar>k\<bar> div \<bar>l\<bar>"
+proof (cases "l = 0")
+ case True
+ then show ?thesis
+ by simp
+next
+ case False
+ with assms have "(sgn k * \<bar>k\<bar>) div (sgn l * \<bar>l\<bar>) = \<bar>k\<bar> div \<bar>l\<bar>"
+ by (simp add: div_sgn_abs_cancel)
+ then show ?thesis
+ by (simp add: sgn_mult_abs)
+qed
+
+lemma div_dvd_sgn_abs:
+ fixes k l :: int
+ assumes "l dvd k"
+ shows "k div l = (sgn k * sgn l) * (\<bar>k\<bar> div \<bar>l\<bar>)"
+proof (cases "k = 0")
+ case True
+ then show ?thesis
+ by simp
+next
+ case False
+ show ?thesis
+ proof (cases "sgn l = sgn k")
+ case True
+ then show ?thesis
+ by (simp add: div_eq_sgn_abs)
+ next
+ case False
+ with \<open>k \<noteq> 0\<close> assms show ?thesis
+ unfolding divide_int_def [of k l]
+ by (auto simp add: zdiv_int)
+ qed
+qed
+
+lemma div_noneq_sgn_abs:
+ fixes k l :: int
+ assumes "l \<noteq> 0"
+ assumes "sgn k \<noteq> sgn l"
+ shows "k div l = - (\<bar>k\<bar> div \<bar>l\<bar>) - of_bool (\<not> l dvd k)"
+ using assms
+ by (simp only: divide_int_def [of k l], auto simp add: not_less zdiv_int)
+
+lemma sgn_mod:
+ fixes k l :: int
+ assumes "l \<noteq> 0" "\<not> l dvd k"
+ shows "sgn (k mod l) = sgn l"
+proof -
+ from \<open>\<not> l dvd k\<close>
+ have "k mod l \<noteq> 0"
+ by (simp add: dvd_eq_mod_eq_0)
+ show ?thesis
+ using \<open>l \<noteq> 0\<close> \<open>\<not> l dvd k\<close>
+ unfolding modulo_int_def [of k l]
+ by (auto simp add: sgn_1_pos sgn_1_neg mod_greater_zero_iff_not_dvd nat_dvd_iff not_less
+ zless_nat_eq_int_zless [symmetric] elim: nonpos_int_cases)
+qed
+
+lemma abs_mod_less:
+ fixes k l :: int
+ assumes "l \<noteq> 0"
+ shows "\<bar>k mod l\<bar> < \<bar>l\<bar>"
+ using assms unfolding modulo_int_def [of k l]
+ by (auto simp add: not_less int_dvd_iff mod_greater_zero_iff_not_dvd elim: pos_int_cases neg_int_cases nonneg_int_cases nonpos_int_cases)
+
+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 "eucl_rel_int (k + s * l) l (s + k div l, k mod l)"
+ using eucl_rel_int [of k l]
+ unfolding eucl_rel_int_iff 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. eucl_rel_int k l (q, r)
+ \<Longrightarrow> eucl_rel_int (s * k) (s * l) (q, s * r)"
+ unfolding eucl_rel_int_iff
+ 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 "eucl_rel_int (s * k) (s * l) (k div l, s * (k mod l))"
+ using eucl_rel_int [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,23 +1956,27 @@
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 pos_mod_conj: "(0::int) < b \<Longrightarrow> 0 \<le> a mod b \<and> a mod b < b"
- using divmod_int_rel [of a b]
- by (auto simp add: divmod_int_rel_def prod_eq_iff)
+ using eucl_rel_int [of a b]
+ by (auto simp add: eucl_rel_int_iff prod_eq_iff)
lemmas pos_mod_sign [simp] = pos_mod_conj [THEN conjunct1]
and pos_mod_bound [simp] = pos_mod_conj [THEN conjunct2]
lemma neg_mod_conj: "b < (0::int) \<Longrightarrow> a mod b \<le> 0 \<and> b < a mod b"
- using divmod_int_rel [of a b]
- by (auto simp add: divmod_int_rel_def prod_eq_iff)
+ using eucl_rel_int [of a b]
+ by (auto simp add: eucl_rel_int_iff prod_eq_iff)
lemmas neg_mod_sign [simp] = neg_mod_conj [THEN conjunct1]
and neg_mod_bound [simp] = neg_mod_conj [THEN conjunct2]
@@ -1833,34 +1986,34 @@
lemma div_pos_pos_trivial: "[| (0::int) \<le> a; a < b |] ==> a div b = 0"
apply (rule div_int_unique)
-apply (auto simp add: divmod_int_rel_def)
+apply (auto simp add: eucl_rel_int_iff)
done
lemma div_neg_neg_trivial: "[| a \<le> (0::int); b < a |] ==> a div b = 0"
apply (rule div_int_unique)
-apply (auto simp add: divmod_int_rel_def)
+apply (auto simp add: eucl_rel_int_iff)
done
lemma div_pos_neg_trivial: "[| (0::int) < a; a+b \<le> 0 |] ==> a div b = -1"
apply (rule div_int_unique)
-apply (auto simp add: divmod_int_rel_def)
+apply (auto simp add: eucl_rel_int_iff)
done
(*There is no div_neg_pos_trivial because 0 div b = 0 would supersede it*)
lemma mod_pos_pos_trivial: "[| (0::int) \<le> a; a < b |] ==> a mod b = a"
apply (rule_tac q = 0 in mod_int_unique)
-apply (auto simp add: divmod_int_rel_def)
+apply (auto simp add: eucl_rel_int_iff)
done
lemma mod_neg_neg_trivial: "[| a \<le> (0::int); b < a |] ==> a mod b = a"
apply (rule_tac q = 0 in mod_int_unique)
-apply (auto simp add: divmod_int_rel_def)
+apply (auto simp add: eucl_rel_int_iff)
done
lemma mod_pos_neg_trivial: "[| (0::int) < a; a+b \<le> 0 |] ==> a mod b = a+b"
apply (rule_tac q = "-1" in mod_int_unique)
-apply (auto simp add: divmod_int_rel_def)
+apply (auto simp add: eucl_rel_int_iff)
done
text\<open>There is no \<open>mod_neg_pos_trivial\<close>.\<close>
@@ -1869,28 +2022,33 @@
subsubsection \<open>Laws for div and mod with Unary Minus\<close>
lemma zminus1_lemma:
- "divmod_int_rel a b (q, r) ==> b \<noteq> 0
- ==> divmod_int_rel (-a) b (if r=0 then -q else -q - 1,
+ "eucl_rel_int a b (q, r) ==> b \<noteq> 0
+ ==> eucl_rel_int (-a) b (if r=0 then -q else -q - 1,
if r=0 then 0 else b-r)"
-by (force simp add: split_ifs divmod_int_rel_def linorder_neq_iff right_diff_distrib)
+by (force simp add: split_ifs eucl_rel_int_iff linorder_neq_iff right_diff_distrib)
lemma zdiv_zminus1_eq_if:
"b \<noteq> (0::int)
==> (-a) div b =
(if a mod b = 0 then - (a div b) else - (a div b) - 1)"
-by (blast intro: divmod_int_rel [THEN zminus1_lemma, THEN div_int_unique])
+by (blast intro: eucl_rel_int [THEN zminus1_lemma, THEN div_int_unique])
lemma zmod_zminus1_eq_if:
"(-a::int) mod b = (if a mod b = 0 then 0 else b - (a mod b))"
apply (case_tac "b = 0", simp)
-apply (blast intro: divmod_int_rel [THEN zminus1_lemma, THEN mod_int_unique])
+apply (blast intro: eucl_rel_int [THEN zminus1_lemma, THEN mod_int_unique])
done
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 +2060,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>
@@ -1998,27 +2151,27 @@
text\<open>proving (a*b) div c = a * (b div c) + a * (b mod c)\<close>
lemma zmult1_lemma:
- "[| divmod_int_rel b c (q, r) |]
- ==> divmod_int_rel (a * b) c (a*q + a*r div c, a*r mod c)"
-by (auto simp add: split_ifs divmod_int_rel_def linorder_neq_iff distrib_left ac_simps)
+ "[| eucl_rel_int b c (q, r) |]
+ ==> eucl_rel_int (a * b) c (a*q + a*r div c, a*r mod c)"
+by (auto simp add: split_ifs eucl_rel_int_iff linorder_neq_iff distrib_left ac_simps)
lemma zdiv_zmult1_eq: "(a*b) div c = a*(b div c) + a*(b mod c) div (c::int)"
apply (case_tac "c = 0", simp)
-apply (blast intro: divmod_int_rel [THEN zmult1_lemma, THEN div_int_unique])
+apply (blast intro: eucl_rel_int [THEN zmult1_lemma, THEN div_int_unique])
done
text\<open>proving (a+b) div c = a div c + b div c + ((a mod c + b mod c) div c)\<close>
lemma zadd1_lemma:
- "[| divmod_int_rel a c (aq, ar); divmod_int_rel b c (bq, br) |]
- ==> divmod_int_rel (a+b) c (aq + bq + (ar+br) div c, (ar+br) mod c)"
-by (force simp add: split_ifs divmod_int_rel_def linorder_neq_iff distrib_left)
+ "[| eucl_rel_int a c (aq, ar); eucl_rel_int b c (bq, br) |]
+ ==> eucl_rel_int (a+b) c (aq + bq + (ar+br) div c, (ar+br) mod c)"
+by (force simp add: split_ifs eucl_rel_int_iff linorder_neq_iff distrib_left)
(*NOT suitable for rewriting: the RHS has an instance of the LHS*)
lemma zdiv_zadd1_eq:
"(a+b) div (c::int) = a div c + b div c + ((a mod c + b mod c) div c)"
apply (case_tac "c = 0", simp)
-apply (blast intro: zadd1_lemma [OF divmod_int_rel divmod_int_rel] div_int_unique)
+apply (blast intro: zadd1_lemma [OF eucl_rel_int eucl_rel_int] div_int_unique)
done
lemma zmod_eq_0_iff: "(m mod d = 0) = (EX q::int. m = d*q)"
@@ -2071,9 +2224,9 @@
apply simp
done
-lemma zmult2_lemma: "[| divmod_int_rel a b (q, r); 0 < c |]
- ==> divmod_int_rel a (b * c) (q div c, b*(q mod c) + r)"
-by (auto simp add: mult.assoc divmod_int_rel_def linorder_neq_iff
+lemma zmult2_lemma: "[| eucl_rel_int a b (q, r); 0 < c |]
+ ==> eucl_rel_int a (b * c) (q div c, b*(q mod c) + r)"
+by (auto simp add: mult.assoc eucl_rel_int_iff linorder_neq_iff
zero_less_mult_iff distrib_left [symmetric]
zmult2_lemma_aux1 zmult2_lemma_aux2 zmult2_lemma_aux3 zmult2_lemma_aux4 mult_less_0_iff split: if_split_asm)
@@ -2081,14 +2234,14 @@
fixes a b c :: int
shows "0 \<le> c \<Longrightarrow> a div (b * c) = (a div b) div c"
apply (case_tac "b = 0", simp)
-apply (force simp add: le_less divmod_int_rel [THEN zmult2_lemma, THEN div_int_unique])
+apply (force simp add: le_less eucl_rel_int [THEN zmult2_lemma, THEN div_int_unique])
done
lemma zmod_zmult2_eq:
fixes a b c :: int
shows "0 \<le> c \<Longrightarrow> a mod (b * c) = b * (a div b mod c) + a mod b"
apply (case_tac "b = 0", simp)
-apply (force simp add: le_less divmod_int_rel [THEN zmult2_lemma, THEN mod_int_unique])
+apply (force simp add: le_less eucl_rel_int [THEN zmult2_lemma, THEN mod_int_unique])
done
lemma div_pos_geq:
@@ -2121,7 +2274,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 +2287,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>
@@ -2175,27 +2328,27 @@
subsubsection \<open>Computing \<open>div\<close> and \<open>mod\<close> with shifting\<close>
-lemma pos_divmod_int_rel_mult_2:
+lemma pos_eucl_rel_int_mult_2:
assumes "0 \<le> b"
- assumes "divmod_int_rel a b (q, r)"
- shows "divmod_int_rel (1 + 2*a) (2*b) (q, 1 + 2*r)"
- using assms unfolding divmod_int_rel_def by auto
-
-lemma neg_divmod_int_rel_mult_2:
+ assumes "eucl_rel_int a b (q, r)"
+ shows "eucl_rel_int (1 + 2*a) (2*b) (q, 1 + 2*r)"
+ using assms unfolding eucl_rel_int_iff by auto
+
+lemma neg_eucl_rel_int_mult_2:
assumes "b \<le> 0"
- assumes "divmod_int_rel (a + 1) b (q, r)"
- shows "divmod_int_rel (1 + 2*a) (2*b) (q, 2*r - 1)"
- using assms unfolding divmod_int_rel_def by auto
+ assumes "eucl_rel_int (a + 1) b (q, r)"
+ shows "eucl_rel_int (1 + 2*a) (2*b) (q, 2*r - 1)"
+ using assms unfolding eucl_rel_int_iff by auto
text\<open>computing div by shifting\<close>
lemma pos_zdiv_mult_2: "(0::int) \<le> a ==> (1 + 2*b) div (2*a) = b div a"
- using pos_divmod_int_rel_mult_2 [OF _ divmod_int_rel]
+ using pos_eucl_rel_int_mult_2 [OF _ eucl_rel_int]
by (rule div_int_unique)
lemma neg_zdiv_mult_2:
assumes A: "a \<le> (0::int)" shows "(1 + 2*b) div (2*a) = (b+1) div a"
- using neg_divmod_int_rel_mult_2 [OF A divmod_int_rel]
+ using neg_eucl_rel_int_mult_2 [OF A eucl_rel_int]
by (rule div_int_unique)
(* FIXME: add rules for negative numerals *)
@@ -2216,14 +2369,14 @@
fixes a b :: int
assumes "0 \<le> a"
shows "(1 + 2 * b) mod (2 * a) = 1 + 2 * (b mod a)"
- using pos_divmod_int_rel_mult_2 [OF assms divmod_int_rel]
+ using pos_eucl_rel_int_mult_2 [OF assms eucl_rel_int]
by (rule mod_int_unique)
lemma neg_zmod_mult_2:
fixes a b :: int
assumes "a \<le> 0"
shows "(1 + 2 * b) mod (2 * a) = 2 * ((b + 1) mod a) - 1"
- using neg_divmod_int_rel_mult_2 [OF assms divmod_int_rel]
+ using neg_eucl_rel_int_mult_2 [OF assms eucl_rel_int]
by (rule mod_int_unique)
(* FIXME: add rules for negative numerals *)
@@ -2428,19 +2581,19 @@
text \<open>Simplify expresions in which div and mod combine numerical constants\<close>
lemma int_div_pos_eq: "\<lbrakk>(a::int) = b * q + r; 0 \<le> r; r < b\<rbrakk> \<Longrightarrow> a div b = q"
- by (rule div_int_unique [of a b q r]) (simp add: divmod_int_rel_def)
+ by (rule div_int_unique [of a b q r]) (simp add: eucl_rel_int_iff)
lemma int_div_neg_eq: "\<lbrakk>(a::int) = b * q + r; r \<le> 0; b < r\<rbrakk> \<Longrightarrow> a div b = q"
by (rule div_int_unique [of a b q r],
- simp add: divmod_int_rel_def)
+ simp add: eucl_rel_int_iff)
lemma int_mod_pos_eq: "\<lbrakk>(a::int) = b * q + r; 0 \<le> r; r < b\<rbrakk> \<Longrightarrow> a mod b = r"
by (rule mod_int_unique [of a b q r],
- simp add: divmod_int_rel_def)
+ simp add: eucl_rel_int_iff)
lemma int_mod_neg_eq: "\<lbrakk>(a::int) = b * q + r; r \<le> 0; b < r\<rbrakk> \<Longrightarrow> a mod b = r"
by (rule mod_int_unique [of a b q r],
- simp add: divmod_int_rel_def)
+ simp add: eucl_rel_int_iff)
lemma abs_div: "(y::int) dvd x \<Longrightarrow> \<bar>x div y\<bar> = \<bar>x\<bar> div \<bar>y\<bar>"
by (unfold dvd_def, cases "y=0", auto simp add: abs_mult)
@@ -2454,24 +2607,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 +2666,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 +2802,4 @@
declare minus_div_mult_eq_mod [symmetric, nitpick_unfold]
-hide_fact (open) div_0 div_by_0
-
end
--- a/src/HOL/Enum.thy Wed Jan 04 14:26:08 2017 +0100
+++ b/src/HOL/Enum.thy Wed Jan 04 14:26:19 2017 +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 Wed Jan 04 14:26:08 2017 +0100
+++ b/src/HOL/Fields.thy Wed Jan 04 14:26:19 2017 +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 Wed Jan 04 14:26:08 2017 +0100
+++ b/src/HOL/Fun_Def.thy Wed Jan 04 14:26:19 2017 +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 Wed Jan 04 14:26:08 2017 +0100
+++ b/src/HOL/GCD.thy Wed Jan 04 14:26:19 2017 +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 Wed Jan 04 14:26:08 2017 +0100
+++ b/src/HOL/Groebner_Basis.thy Wed Jan 04 14:26:19 2017 +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 Wed Jan 04 14:26:08 2017 +0100
+++ b/src/HOL/Hilbert_Choice.thy Wed Jan 04 14:26:19 2017 +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 Wed Jan 04 14:26:08 2017 +0100
+++ b/src/HOL/IMP/Abs_Int1_parity.thy Wed Jan 04 14:26:19 2017 +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/Inductive.thy Wed Jan 04 14:26:08 2017 +0100
+++ b/src/HOL/Inductive.thy Wed Jan 04 14:26:19 2017 +0100
@@ -519,9 +519,8 @@
ML_file "Tools/Old_Datatype/old_datatype_data.ML"
ML_file "Tools/Old_Datatype/old_rep_datatype.ML"
ML_file "Tools/Old_Datatype/old_datatype_codegen.ML"
+ML_file "Tools/BNF/bnf_fp_rec_sugar_util.ML"
ML_file "Tools/Old_Datatype/old_primrec.ML"
-
-ML_file "Tools/BNF/bnf_fp_rec_sugar_util.ML"
ML_file "Tools/BNF/bnf_lfp_rec_sugar.ML"
text \<open>Lambda-abstractions with pattern matching:\<close>
--- a/src/HOL/Int.thy Wed Jan 04 14:26:08 2017 +0100
+++ b/src/HOL/Int.thy Wed Jan 04 14:26:19 2017 +0100
@@ -433,11 +433,57 @@
lemma zless_nat_conj [simp]: "nat w < nat z \<longleftrightarrow> 0 < z \<and> w < z"
by transfer (clarsimp, arith)
-lemma nonneg_eq_int:
- fixes z :: int
- assumes "0 \<le> z" and "\<And>m. z = int m \<Longrightarrow> P"
- shows P
- using assms by (blast dest: nat_0_le sym)
+lemma nonneg_int_cases:
+ assumes "0 \<le> k"
+ obtains n where "k = int n"
+proof -
+ from assms have "k = int (nat k)"
+ by simp
+ then show thesis
+ by (rule that)
+qed
+
+lemma pos_int_cases:
+ assumes "0 < k"
+ obtains n where "k = int n" and "n > 0"
+proof -
+ from assms have "0 \<le> k"
+ by simp
+ then obtain n where "k = int n"
+ by (rule nonneg_int_cases)
+ moreover have "n > 0"
+ using \<open>k = int n\<close> assms by simp
+ ultimately show thesis
+ by (rule that)
+qed
+
+lemma nonpos_int_cases:
+ assumes "k \<le> 0"
+ obtains n where "k = - int n"
+proof -
+ from assms have "- k \<ge> 0"
+ by simp
+ then obtain n where "- k = int n"
+ by (rule nonneg_int_cases)
+ then have "k = - int n"
+ by simp
+ then show thesis
+ by (rule that)
+qed
+
+lemma neg_int_cases:
+ assumes "k < 0"
+ obtains n where "k = - int n" and "n > 0"
+proof -
+ from assms have "- k > 0"
+ by simp
+ then obtain n where "- k = int n" and "- k > 0"
+ by (blast elim: pos_int_cases)
+ then have "k = - int n" and "n > 0"
+ by simp_all
+ then show thesis
+ by (rule that)
+qed
lemma nat_eq_iff: "nat w = m \<longleftrightarrow> (if 0 \<le> w then w = int m else m = 0)"
by transfer (clarsimp simp add: le_imp_diff_is_add)
@@ -615,11 +661,6 @@
"(\<And>n. P (int n)) \<Longrightarrow> (\<And>n. P (- (int (Suc n)))) \<Longrightarrow> P z"
by (cases z) auto
-lemma nonneg_int_cases:
- assumes "0 \<le> k"
- obtains n where "k = int n"
- using assms by (rule nonneg_eq_int)
-
lemma Let_numeral [simp]: "Let (numeral v) f = f (numeral v)"
\<comment> \<open>Unfold all \<open>let\<close>s involving constants\<close>
by (fact Let_numeral) \<comment> \<open>FIXME drop\<close>
@@ -791,6 +832,10 @@
end
+lemma (in linordered_idom) Ints_abs [simp]:
+ shows "a \<in> \<int> \<Longrightarrow> abs a \<in> \<int>"
+ by (auto simp: abs_if)
+
lemma (in linordered_idom) Nats_altdef2: "\<nat> = {n \<in> \<int>. n \<ge> 0}"
proof (intro subsetI equalityI)
fix x :: 'a
@@ -880,14 +925,14 @@
lemma of_int_prod [simp]: "of_int (prod f A) = (\<Prod>x\<in>A. of_int(f x))"
by (induct A rule: infinite_finite_induct) auto
-lemmas int_sum = of_nat_sum [where 'a=int]
-lemmas int_prod = of_nat_prod [where 'a=int]
-
text \<open>Legacy theorems\<close>
+lemmas int_sum = of_nat_sum [where 'a=int]
+lemmas int_prod = of_nat_prod [where 'a=int]
lemmas zle_int = of_nat_le_iff [where 'a=int]
lemmas int_int_eq = of_nat_eq_iff [where 'a=int]
+lemmas nonneg_eq_int = nonneg_int_cases
subsection \<open>Setting up simplification procedures\<close>
@@ -927,6 +972,24 @@
for z :: int
by arith
+lemma Ints_nonzero_abs_ge1:
+ fixes x:: "'a :: linordered_idom"
+ assumes "x \<in> Ints" "x \<noteq> 0"
+ shows "1 \<le> abs x"
+proof (rule Ints_cases [OF \<open>x \<in> Ints\<close>])
+ fix z::int
+ assume "x = of_int z"
+ with \<open>x \<noteq> 0\<close>
+ show "1 \<le> \<bar>x\<bar>"
+ apply (auto simp add: abs_if)
+ by (metis diff_0 of_int_1 of_int_le_iff of_int_minus zle_diff1_eq)
+qed
+
+lemma Ints_nonzero_abs_less1:
+ fixes x:: "'a :: linordered_idom"
+ shows "\<lbrakk>x \<in> Ints; abs x < 1\<rbrakk> \<Longrightarrow> x = 0"
+ using Ints_nonzero_abs_ge1 [of x] by auto
+
subsection \<open>The functions @{term nat} and @{term int}\<close>
--- a/src/HOL/Library/Code_Test.thy Wed Jan 04 14:26:08 2017 +0100
+++ b/src/HOL/Library/Code_Test.thy Wed Jan 04 14:26:19 2017 +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 Wed Jan 04 14:26:08 2017 +0100
+++ b/src/HOL/Library/DAList_Multiset.thy Wed Jan 04 14:26:19 2017 +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 Wed Jan 04 14:26:19 2017 +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 Wed Jan 04 14:26:08 2017 +0100
+++ b/src/HOL/Library/Formal_Power_Series.thy Wed Jan 04 14:26:19 2017 +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/Infinite_Set.thy Wed Jan 04 14:26:08 2017 +0100
+++ b/src/HOL/Library/Infinite_Set.thy Wed Jan 04 14:26:19 2017 +0100
@@ -152,7 +152,7 @@
by (simp add: eventually_frequently)
lemma MOST_Suc_iff: "(MOST n. P (Suc n)) \<longleftrightarrow> (MOST n. P n)"
- by (simp add: cofinite_eq_sequentially eventually_sequentially_Suc)
+ by (simp add: cofinite_eq_sequentially)
lemma
shows MOST_SucI: "MOST n. P n \<Longrightarrow> MOST n. P (Suc n)"
@@ -317,7 +317,7 @@
proof (induction A)
case (insert a A)
with R show ?case
- by (metis empty_iff insert_iff)
+ by (metis empty_iff insert_iff)
qed simp
corollary Union_maximal_sets:
--- a/src/HOL/Library/Library.thy Wed Jan 04 14:26:08 2017 +0100
+++ b/src/HOL/Library/Library.thy Wed Jan 04 14:26:19 2017 +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 Wed Jan 04 14:26:08 2017 +0100
+++ b/src/HOL/Library/Multiset.thy Wed Jan 04 14:26:19 2017 +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 Wed Jan 04 14:26:08 2017 +0100
+++ b/src/HOL/Library/Multiset_Order.thy Wed Jan 04 14:26:19 2017 +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 Wed Jan 04 14:26:08 2017 +0100
+++ b/src/HOL/Library/Normalized_Fraction.thy Wed Jan 04 14:26:19 2017 +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 Wed Jan 04 14:26:08 2017 +0100
+++ b/src/HOL/Library/Numeral_Type.thy Wed Jan 04 14:26:19 2017 +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/Old_SMT/old_z3_proof_reconstruction.ML Wed Jan 04 14:26:08 2017 +0100
+++ b/src/HOL/Library/Old_SMT/old_z3_proof_reconstruction.ML Wed Jan 04 14:26:19 2017 +0100
@@ -167,7 +167,7 @@
assms
|> map (apsnd (rewrite ctxt eqs'))
|> map (apsnd (Conv.fconv_rule Thm.eta_conversion))
- |> Old_Z3_Proof_Tools.thm_net_of snd
+ |> Old_Z3_Proof_Tools.thm_net_of snd
fun revert_conv ctxt = rewrite_conv ctxt eqs' then_conv Thm.eta_conversion
@@ -183,14 +183,14 @@
if exact then (Thm.implies_elim thm1 th, ctxt)
else assume thm1 ctxt
val thms' = if exact then thms else th :: thms
- in
+ in
((insert (op =) i is, thms'),
(ctxt', Inttab.update (idx, Thm thm) ptab))
end
fun add (idx, ct) (cx as ((is, thms), (ctxt, ptab))) =
let
- val thm1 =
+ val thm1 =
Thm.trivial ct
|> Conv.fconv_rule (Conv.arg1_conv (revert_conv outer_ctxt))
val thm2 = singleton (Variable.export ctxt outer_ctxt) thm1
@@ -218,7 +218,7 @@
val mp_c = precomp (Thm.dest_binop o Thm.dest_arg) @{thm mp}
in
fun mp (MetaEq thm) p = Thm (Thm.implies_elim (comp meta_iffD1_c thm) p)
- | mp p_q p =
+ | mp p_q p =
let
val pq = thm_of p_q
val thm = comp iffD1_c pq handle THM _ => comp mp_c pq
@@ -509,7 +509,7 @@
(case lookup (Logic.mk_equals (apply2 Thm.term_of cp)) of
SOME eq => eq
| NONE => if exn then raise MONO else prove_refl cp)
-
+
val prove_exn = prove_eq true
and prove_safe = prove_eq false
@@ -752,7 +752,9 @@
fun check_after idx r ps ct (p, (ctxt, _)) =
if not (Config.get ctxt Old_SMT_Config.trace) then ()
else
- let val thm = thm_of p |> tap (Thm.join_proofs o single)
+ let
+ val thm = thm_of p
+ val _ = Thm.consolidate thm
in
if (Thm.cprop_of thm) aconvc ct then ()
else
@@ -852,7 +854,7 @@
fun discharge_assms_tac ctxt rules =
REPEAT (HEADGOAL (resolve_tac ctxt rules ORELSE' SOLVED' (discharge_sk_tac ctxt)))
-
+
fun discharge_assms ctxt rules thm =
if Thm.nprems_of thm = 0 then Goal.norm_result ctxt thm
else
@@ -881,7 +883,7 @@
if Config.get ctxt2 Old_SMT_Config.filter_only_facts then (is, @{thm TrueI})
else
(Thm @{thm TrueI}, cxp)
- |> fold (prove simpset vars) steps
+ |> fold (prove simpset vars) steps
|> discharge rules outer_ctxt
|> pair []
end
--- a/src/HOL/Library/Omega_Words_Fun.thy Wed Jan 04 14:26:08 2017 +0100
+++ b/src/HOL/Library/Omega_Words_Fun.thy Wed Jan 04 14:26:19 2017 +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 Wed Jan 04 14:26:08 2017 +0100
+++ b/src/HOL/Library/Permutation.thy Wed Jan 04 14:26:19 2017 +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/Permutations.thy Wed Jan 04 14:26:08 2017 +0100
+++ b/src/HOL/Library/Permutations.thy Wed Jan 04 14:26:19 2017 +0100
@@ -31,7 +31,7 @@
using surj_f_inv_f[OF bij_is_surj[OF bp]]
by (simp add: fun_eq_iff Fun.swap_def bij_inv_eq_iff[OF bp])
-lemma bij_swap_ompose_bij: "bij p \<Longrightarrow> bij (Fun.swap a b id \<circ> p)"
+lemma bij_swap_compose_bij: "bij p \<Longrightarrow> bij (Fun.swap a b id \<circ> p)"
proof -
assume H: "bij p"
show ?thesis
@@ -756,18 +756,10 @@
let ?q = "Fun.swap a (p a) id \<circ> ?r"
have raa: "?r a = a"
by (simp add: Fun.swap_def)
- from bij_swap_ompose_bij[OF insert(4)]
- have br: "bij ?r" .
-
+ from bij_swap_compose_bij[OF insert(4)] have br: "bij ?r" .
from insert raa have th: "\<forall>x. x \<notin> F \<longrightarrow> ?r x = x"
- apply (clarsimp simp add: Fun.swap_def)
- apply (erule_tac x="x" in allE)
- apply auto
- unfolding bij_iff
- apply metis
- done
- from insert(3)[OF br th]
- have rp: "permutation ?r" .
+ by (metis bij_pointE comp_apply id_apply insert_iff swap_apply(3))
+ from insert(3)[OF br th] have rp: "permutation ?r" .
have "permutation ?q"
by (simp add: permutation_compose permutation_swap_id rp)
then show ?case
@@ -926,7 +918,7 @@
using permutes_in_image[OF assms] by auto
have "count (mset (permute_list f xs)) y =
card ((\<lambda>i. xs ! f i) -` {y} \<inter> {..<length xs})"
- by (simp add: permute_list_def mset_map count_image_mset atLeast0LessThan)
+ by (simp add: permute_list_def count_image_mset atLeast0LessThan)
also have "(\<lambda>i. xs ! f i) -` {y} \<inter> {..<length xs} = f -` {i. i < length xs \<and> y = xs ! i}"
by auto
also from assms have "card \<dots> = card {i. i < length xs \<and> y = xs ! i}"
--- a/src/HOL/Library/Polynomial.thy Wed Jan 04 14:26:08 2017 +0100
+++ b/src/HOL/Library/Polynomial.thy Wed Jan 04 14:26:19 2017 +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>
@@ -1433,18 +1613,26 @@
subsection \<open>Long division of polynomials\<close>
-definition pdivmod_rel :: "'a::field poly \<Rightarrow> 'a poly \<Rightarrow> 'a poly \<Rightarrow> 'a poly \<Rightarrow> bool"
-where
- "pdivmod_rel x y q r \<longleftrightarrow>
- x = q * y + r \<and> (if y = 0 then q = 0 else r = 0 \<or> degree r < degree y)"
-
-lemma pdivmod_rel_0:
- "pdivmod_rel 0 y 0 0"
- unfolding pdivmod_rel_def by simp
-
-lemma pdivmod_rel_by_0:
- "pdivmod_rel x 0 0 x"
- unfolding pdivmod_rel_def by simp
+inductive eucl_rel_poly :: "'a::field poly \<Rightarrow> 'a poly \<Rightarrow> 'a poly \<times> 'a poly \<Rightarrow> bool"
+ where eucl_rel_poly_by0: "eucl_rel_poly x 0 (0, x)"
+ | eucl_rel_poly_dividesI: "y \<noteq> 0 \<Longrightarrow> x = q * y \<Longrightarrow> eucl_rel_poly x y (q, 0)"
+ | eucl_rel_poly_remainderI: "y \<noteq> 0 \<Longrightarrow> degree r < degree y
+ \<Longrightarrow> x = q * y + r \<Longrightarrow> eucl_rel_poly x y (q, r)"
+
+lemma eucl_rel_poly_iff:
+ "eucl_rel_poly x y (q, r) \<longleftrightarrow>
+ x = q * y + r \<and>
+ (if y = 0 then q = 0 else r = 0 \<or> degree r < degree y)"
+ by (auto elim: eucl_rel_poly.cases
+ intro: eucl_rel_poly_by0 eucl_rel_poly_dividesI eucl_rel_poly_remainderI)
+
+lemma eucl_rel_poly_0:
+ "eucl_rel_poly 0 y (0, 0)"
+ unfolding eucl_rel_poly_iff by simp
+
+lemma eucl_rel_poly_by_0:
+ "eucl_rel_poly x 0 (0, x)"
+ unfolding eucl_rel_poly_iff by simp
lemma eq_zero_or_degree_less:
assumes "degree p \<le> n" and "coeff p n = 0"
@@ -1470,15 +1658,15 @@
then show ?thesis ..
qed
-lemma pdivmod_rel_pCons:
- assumes rel: "pdivmod_rel x y q r"
+lemma eucl_rel_poly_pCons:
+ assumes rel: "eucl_rel_poly x y (q, r)"
assumes y: "y \<noteq> 0"
assumes b: "b = coeff (pCons a r) (degree y) / coeff y (degree y)"
- shows "pdivmod_rel (pCons a x) y (pCons b q) (pCons a r - smult b y)"
- (is "pdivmod_rel ?x y ?q ?r")
+ shows "eucl_rel_poly (pCons a x) y (pCons b q, pCons a r - smult b y)"
+ (is "eucl_rel_poly ?x y (?q, ?r)")
proof -
have x: "x = q * y + r" and r: "r = 0 \<or> degree r < degree y"
- using assms unfolding pdivmod_rel_def by simp_all
+ using assms unfolding eucl_rel_poly_iff by simp_all
have 1: "?x = ?q * y + ?r"
using b x by simp
@@ -1498,31 +1686,31 @@
qed
from 1 2 show ?thesis
- unfolding pdivmod_rel_def
+ unfolding eucl_rel_poly_iff
using \<open>y \<noteq> 0\<close> by simp
qed
-lemma pdivmod_rel_exists: "\<exists>q r. pdivmod_rel x y q r"
+lemma eucl_rel_poly_exists: "\<exists>q r. eucl_rel_poly x y (q, r)"
apply (cases "y = 0")
-apply (fast intro!: pdivmod_rel_by_0)
+apply (fast intro!: eucl_rel_poly_by_0)
apply (induct x)
-apply (fast intro!: pdivmod_rel_0)
-apply (fast intro!: pdivmod_rel_pCons)
+apply (fast intro!: eucl_rel_poly_0)
+apply (fast intro!: eucl_rel_poly_pCons)
done
-lemma pdivmod_rel_unique:
- assumes 1: "pdivmod_rel x y q1 r1"
- assumes 2: "pdivmod_rel x y q2 r2"
+lemma eucl_rel_poly_unique:
+ assumes 1: "eucl_rel_poly x y (q1, r1)"
+ assumes 2: "eucl_rel_poly x y (q2, r2)"
shows "q1 = q2 \<and> r1 = r2"
proof (cases "y = 0")
assume "y = 0" with assms show ?thesis
- by (simp add: pdivmod_rel_def)
+ by (simp add: eucl_rel_poly_iff)
next
assume [simp]: "y \<noteq> 0"
from 1 have q1: "x = q1 * y + r1" and r1: "r1 = 0 \<or> degree r1 < degree y"
- unfolding pdivmod_rel_def by simp_all
+ unfolding eucl_rel_poly_iff by simp_all
from 2 have q2: "x = q2 * y + r2" and r2: "r2 = 0 \<or> degree r2 < degree y"
- unfolding pdivmod_rel_def by simp_all
+ unfolding eucl_rel_poly_iff by simp_all
from q1 q2 have q3: "(q1 - q2) * y = r2 - r1"
by (simp add: algebra_simps)
from r1 r2 have r3: "(r2 - r1) = 0 \<or> degree (r2 - r1) < degree y"
@@ -1543,19 +1731,19 @@
qed
qed
-lemma pdivmod_rel_0_iff: "pdivmod_rel 0 y q r \<longleftrightarrow> q = 0 \<and> r = 0"
-by (auto dest: pdivmod_rel_unique intro: pdivmod_rel_0)
-
-lemma pdivmod_rel_by_0_iff: "pdivmod_rel x 0 q r \<longleftrightarrow> q = 0 \<and> r = x"
-by (auto dest: pdivmod_rel_unique intro: pdivmod_rel_by_0)
-
-lemmas pdivmod_rel_unique_div = pdivmod_rel_unique [THEN conjunct1]
-
-lemmas pdivmod_rel_unique_mod = pdivmod_rel_unique [THEN conjunct2]
-
-
-
-subsection\<open>Pseudo-Division and Division of Polynomials\<close>
+lemma eucl_rel_poly_0_iff: "eucl_rel_poly 0 y (q, r) \<longleftrightarrow> q = 0 \<and> r = 0"
+by (auto dest: eucl_rel_poly_unique intro: eucl_rel_poly_0)
+
+lemma eucl_rel_poly_by_0_iff: "eucl_rel_poly x 0 (q, r) \<longleftrightarrow> q = 0 \<and> r = x"
+by (auto dest: eucl_rel_poly_unique intro: eucl_rel_poly_by_0)
+
+lemmas eucl_rel_poly_unique_div = eucl_rel_poly_unique [THEN conjunct1]
+
+lemmas eucl_rel_poly_unique_mod = eucl_rel_poly_unique [THEN conjunct2]
+
+
+
+subsection \<open>Pseudo-Division and Division of Polynomials\<close>
text\<open>This part is by René Thiemann and Akihisa Yamada.\<close>
@@ -1838,15 +2026,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.
@@ -1899,8 +2244,8 @@
if g = 0 then f
else pseudo_mod (smult ((1/coeff g (degree g)) ^ (Suc (degree f) - degree g)) f) g"
-lemma pdivmod_rel: "pdivmod_rel (x::'a::field poly) y (x div y) (x mod y)"
- unfolding pdivmod_rel_def
+lemma eucl_rel_poly: "eucl_rel_poly (x::'a::field poly) y (x div y, x mod y)"
+ unfolding eucl_rel_poly_iff
proof (intro conjI)
show "x = x div y * y + x mod y"
proof(cases "y = 0")
@@ -1924,30 +2269,24 @@
qed
lemma div_poly_eq:
- "pdivmod_rel (x::'a::field poly) y q r \<Longrightarrow> x div y = q"
- by(rule pdivmod_rel_unique_div[OF pdivmod_rel])
+ "eucl_rel_poly (x::'a::field poly) y (q, r) \<Longrightarrow> x div y = q"
+ by(rule eucl_rel_poly_unique_div[OF eucl_rel_poly])
lemma mod_poly_eq:
- "pdivmod_rel (x::'a::field poly) y q r \<Longrightarrow> x mod y = r"
- by (rule pdivmod_rel_unique_mod[OF pdivmod_rel])
+ "eucl_rel_poly (x::'a::field poly) y (q, r) \<Longrightarrow> x mod y = r"
+ by (rule eucl_rel_poly_unique_mod[OF eucl_rel_poly])
instance
proof
fix x y :: "'a poly"
- from pdivmod_rel[of x y,unfolded pdivmod_rel_def]
+ from eucl_rel_poly[of x y,unfolded eucl_rel_poly_iff]
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)"
- using pdivmod_rel [of x y]
- by (simp add: pdivmod_rel_def distrib_right)
+ hence "eucl_rel_poly (x + z * y) y (z + x div y, x mod y)"
+ using eucl_rel_poly [of x y]
+ by (simp add: eucl_rel_poly_iff distrib_right)
thus "(x + z * y) div y = z + x div y"
by (rule div_poly_eq)
next
@@ -1955,112 +2294,60 @@
assume "x \<noteq> 0"
show "(x * y) div (x * z) = y div z"
proof (cases "y \<noteq> 0 \<and> z \<noteq> 0")
- have "\<And>x::'a poly. pdivmod_rel x 0 0 x"
- by (rule pdivmod_rel_by_0)
+ have "\<And>x::'a poly. eucl_rel_poly x 0 (0, x)"
+ by (rule eucl_rel_poly_by_0)
then have [simp]: "\<And>x::'a poly. x div 0 = 0"
by (rule div_poly_eq)
- have "\<And>x::'a poly. pdivmod_rel 0 x 0 0"
- by (rule pdivmod_rel_0)
+ have "\<And>x::'a poly. eucl_rel_poly 0 x (0, 0)"
+ by (rule eucl_rel_poly_0)
then have [simp]: "\<And>x::'a poly. 0 div x = 0"
by (rule div_poly_eq)
case False then show ?thesis by auto
next
case True then have "y \<noteq> 0" and "z \<noteq> 0" by auto
with \<open>x \<noteq> 0\<close>
- have "\<And>q r. pdivmod_rel y z q r \<Longrightarrow> pdivmod_rel (x * y) (x * z) q (x * r)"
- by (auto simp add: pdivmod_rel_def algebra_simps)
+ have "\<And>q r. eucl_rel_poly y z (q, r) \<Longrightarrow> eucl_rel_poly (x * y) (x * z) (q, x * r)"
+ by (auto simp add: eucl_rel_poly_iff algebra_simps)
(rule classical, simp add: degree_mult_eq)
- moreover from pdivmod_rel have "pdivmod_rel y z (y div z) (y mod z)" .
- ultimately have "pdivmod_rel (x * y) (x * z) (y div z) (x * (y mod z))" .
+ moreover from eucl_rel_poly have "eucl_rel_poly y z (y div z, y mod z)" .
+ ultimately have "eucl_rel_poly (x * y) (x * z) (y div z, x * (y mod z))" .
then show ?thesis by (simp add: div_poly_eq)
qed
qed
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]
- unfolding pdivmod_rel_def by simp
+ using eucl_rel_poly [of x y]
+ unfolding eucl_rel_poly_iff by simp
lemma div_poly_less: "degree (x::'a::field poly) < degree y \<Longrightarrow> x div y = 0"
proof -
assume "degree x < degree y"
- hence "pdivmod_rel x y 0 x"
- by (simp add: pdivmod_rel_def)
+ hence "eucl_rel_poly x y (0, x)"
+ by (simp add: eucl_rel_poly_iff)
thus "x div y = 0" by (rule div_poly_eq)
qed
lemma mod_poly_less: "degree x < degree y \<Longrightarrow> x mod y = x"
proof -
assume "degree x < degree y"
- hence "pdivmod_rel x y 0 x"
- by (simp add: pdivmod_rel_def)
+ hence "eucl_rel_poly x y (0, x)"
+ by (simp add: eucl_rel_poly_iff)
thus "x mod y = x" by (rule mod_poly_eq)
qed
-lemma pdivmod_rel_smult_left:
- "pdivmod_rel x y q r
- \<Longrightarrow> pdivmod_rel (smult a x) y (smult a q) (smult a r)"
- unfolding pdivmod_rel_def by (simp add: smult_add_right)
+lemma eucl_rel_poly_smult_left:
+ "eucl_rel_poly x y (q, r)
+ \<Longrightarrow> eucl_rel_poly (smult a x) y (smult a q, smult a r)"
+ unfolding eucl_rel_poly_iff by (simp add: smult_add_right)
lemma div_smult_left: "(smult (a::'a::field) x) div y = smult a (x div y)"
- by (rule div_poly_eq, rule pdivmod_rel_smult_left, rule pdivmod_rel)
+ by (rule div_poly_eq, rule eucl_rel_poly_smult_left, rule eucl_rel_poly)
lemma mod_smult_left: "(smult a x) mod y = smult a (x mod y)"
- by (rule mod_poly_eq, rule pdivmod_rel_smult_left, rule pdivmod_rel)
+ by (rule mod_poly_eq, rule eucl_rel_poly_smult_left, rule eucl_rel_poly)
lemma poly_div_minus_left [simp]:
fixes x y :: "'a::field poly"
@@ -2072,23 +2359,23 @@
shows "(- x) mod y = - (x mod y)"
using mod_smult_left [of "- 1::'a"] by simp
-lemma pdivmod_rel_add_left:
- assumes "pdivmod_rel x y q r"
- assumes "pdivmod_rel x' y q' r'"
- shows "pdivmod_rel (x + x') y (q + q') (r + r')"
- using assms unfolding pdivmod_rel_def
+lemma eucl_rel_poly_add_left:
+ assumes "eucl_rel_poly x y (q, r)"
+ assumes "eucl_rel_poly x' y (q', r')"
+ shows "eucl_rel_poly (x + x') y (q + q', r + r')"
+ using assms unfolding eucl_rel_poly_iff
by (auto simp add: algebra_simps degree_add_less)
lemma poly_div_add_left:
fixes x y z :: "'a::field poly"
shows "(x + y) div z = x div z + y div z"
- using pdivmod_rel_add_left [OF pdivmod_rel pdivmod_rel]
+ using eucl_rel_poly_add_left [OF eucl_rel_poly eucl_rel_poly]
by (rule div_poly_eq)
lemma poly_mod_add_left:
fixes x y z :: "'a::field poly"
shows "(x + y) mod z = x mod z + y mod z"
- using pdivmod_rel_add_left [OF pdivmod_rel pdivmod_rel]
+ using eucl_rel_poly_add_left [OF eucl_rel_poly eucl_rel_poly]
by (rule mod_poly_eq)
lemma poly_div_diff_left:
@@ -2101,17 +2388,17 @@
shows "(x - y) mod z = x mod z - y mod z"
by (simp only: diff_conv_add_uminus poly_mod_add_left poly_mod_minus_left)
-lemma pdivmod_rel_smult_right:
- "\<lbrakk>a \<noteq> 0; pdivmod_rel x y q r\<rbrakk>
- \<Longrightarrow> pdivmod_rel x (smult a y) (smult (inverse a) q) r"
- unfolding pdivmod_rel_def by simp
+lemma eucl_rel_poly_smult_right:
+ "a \<noteq> 0 \<Longrightarrow> eucl_rel_poly x y (q, r)
+ \<Longrightarrow> eucl_rel_poly x (smult a y) (smult (inverse a) q, r)"
+ unfolding eucl_rel_poly_iff by simp
lemma div_smult_right:
"(a::'a::field) \<noteq> 0 \<Longrightarrow> x div (smult a y) = smult (inverse a) (x div y)"
- by (rule div_poly_eq, erule pdivmod_rel_smult_right, rule pdivmod_rel)
+ by (rule div_poly_eq, erule eucl_rel_poly_smult_right, rule eucl_rel_poly)
lemma mod_smult_right: "a \<noteq> 0 \<Longrightarrow> x mod (smult a y) = x mod y"
- by (rule mod_poly_eq, erule pdivmod_rel_smult_right, rule pdivmod_rel)
+ by (rule mod_poly_eq, erule eucl_rel_poly_smult_right, rule eucl_rel_poly)
lemma poly_div_minus_right [simp]:
fixes x y :: "'a::field poly"
@@ -2123,30 +2410,30 @@
shows "x mod (- y) = x mod y"
using mod_smult_right [of "- 1::'a"] by simp
-lemma pdivmod_rel_mult:
- "\<lbrakk>pdivmod_rel x y q r; pdivmod_rel q z q' r'\<rbrakk>
- \<Longrightarrow> pdivmod_rel x (y * z) q' (y * r' + r)"
-apply (cases "z = 0", simp add: pdivmod_rel_def)
-apply (cases "y = 0", simp add: pdivmod_rel_by_0_iff pdivmod_rel_0_iff)
+lemma eucl_rel_poly_mult:
+ "eucl_rel_poly x y (q, r) \<Longrightarrow> eucl_rel_poly q z (q', r')
+ \<Longrightarrow> eucl_rel_poly x (y * z) (q', y * r' + r)"
+apply (cases "z = 0", simp add: eucl_rel_poly_iff)
+apply (cases "y = 0", simp add: eucl_rel_poly_by_0_iff eucl_rel_poly_0_iff)
apply (cases "r = 0")
apply (cases "r' = 0")
-apply (simp add: pdivmod_rel_def)
-apply (simp add: pdivmod_rel_def field_simps degree_mult_eq)
+apply (simp add: eucl_rel_poly_iff)
+apply (simp add: eucl_rel_poly_iff field_simps degree_mult_eq)
apply (cases "r' = 0")
-apply (simp add: pdivmod_rel_def degree_mult_eq)
-apply (simp add: pdivmod_rel_def field_simps)
+apply (simp add: eucl_rel_poly_iff degree_mult_eq)
+apply (simp add: eucl_rel_poly_iff field_simps)
apply (simp add: degree_mult_eq degree_add_less)
done
lemma poly_div_mult_right:
fixes x y z :: "'a::field poly"
shows "x div (y * z) = (x div y) div z"
- by (rule div_poly_eq, rule pdivmod_rel_mult, (rule pdivmod_rel)+)
+ by (rule div_poly_eq, rule eucl_rel_poly_mult, (rule eucl_rel_poly)+)
lemma poly_mod_mult_right:
fixes x y z :: "'a::field poly"
shows "x mod (y * z) = y * (x div y mod z) + x mod y"
- by (rule mod_poly_eq, rule pdivmod_rel_mult, (rule pdivmod_rel)+)
+ by (rule mod_poly_eq, rule eucl_rel_poly_mult, (rule eucl_rel_poly)+)
lemma mod_pCons:
fixes a and x
@@ -2155,7 +2442,7 @@
shows "(pCons a x) mod y = (pCons a (x mod y) - smult b y)"
unfolding b
apply (rule mod_poly_eq)
-apply (rule pdivmod_rel_pCons [OF pdivmod_rel y refl])
+apply (rule eucl_rel_poly_pCons [OF eucl_rel_poly y refl])
done
definition pdivmod :: "'a::field poly \<Rightarrow> 'a poly \<Rightarrow> 'a poly \<times> 'a poly"
@@ -2174,9 +2461,9 @@
in (pCons b s, pCons a r - smult b q)))"
apply (simp add: pdivmod_def Let_def, safe)
apply (rule div_poly_eq)
- apply (erule pdivmod_rel_pCons [OF pdivmod_rel _ refl])
+ apply (erule eucl_rel_poly_pCons [OF eucl_rel_poly _ refl])
apply (rule mod_poly_eq)
- apply (erule pdivmod_rel_pCons [OF pdivmod_rel _ refl])
+ apply (erule eucl_rel_poly_pCons [OF eucl_rel_poly _ refl])
done
lemma pdivmod_fold_coeffs:
@@ -2421,8 +2708,8 @@
(* *************** *)
subsubsection \<open>Improved Code-Equations for Polynomial (Pseudo) Division\<close>
-lemma pdivmod_pdivmodrel: "pdivmod_rel p q r s = (pdivmod p q = (r, s))"
- by (metis pdivmod_def pdivmod_rel pdivmod_rel_unique prod.sel)
+lemma pdivmod_pdivmodrel: "eucl_rel_poly p q (r, s) \<longleftrightarrow> pdivmod p q = (r, s)"
+ by (metis pdivmod_def eucl_rel_poly eucl_rel_poly_unique)
lemma pdivmod_via_pseudo_divmod: "pdivmod f g = (if g = 0 then (0,f)
else let
@@ -2441,7 +2728,7 @@
unfolding Let_def h_def[symmetric] lc_def[symmetric] p by auto
from pseudo_divmod[OF h0 p, unfolded h1]
have f: "f = h * q + r" and r: "r = 0 \<or> degree r < degree h" by auto
- have "pdivmod_rel f h q r" unfolding pdivmod_rel_def using f r h0 by auto
+ have "eucl_rel_poly f h (q, r)" unfolding eucl_rel_poly_iff using f r h0 by auto
hence "pdivmod f h = (q,r)" by (simp add: pdivmod_pdivmodrel)
hence "pdivmod f g = (smult lc q, r)"
unfolding pdivmod_def h_def div_smult_right[OF lc] mod_smult_right[OF lc]
@@ -2860,18 +3147,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 +3212,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 +3242,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 Wed Jan 04 14:26:08 2017 +0100
+++ b/src/HOL/Library/Polynomial_Factorial.thy Wed Jan 04 14:26:19 2017 +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/Transitive_Closure_Table.thy Wed Jan 04 14:26:08 2017 +0100
+++ b/src/HOL/Library/Transitive_Closure_Table.thy Wed Jan 04 14:26:19 2017 +0100
@@ -207,10 +207,10 @@
code_pred rtranclp
using rtranclp_eq_rtrancl_tab_nil [THEN iffD1] by fastforce
-lemma rtrancl_path_Range: "\<lbrakk> rtrancl_path R x xs y; z \<in> set xs \<rbrakk> \<Longrightarrow> RangeP R z"
+lemma rtrancl_path_Range: "\<lbrakk> rtrancl_path R x xs y; z \<in> set xs \<rbrakk> \<Longrightarrow> Rangep R z"
by(induction rule: rtrancl_path.induct) auto
-lemma rtrancl_path_Range_end: "\<lbrakk> rtrancl_path R x xs y; xs \<noteq> [] \<rbrakk> \<Longrightarrow> RangeP R y"
+lemma rtrancl_path_Range_end: "\<lbrakk> rtrancl_path R x xs y; xs \<noteq> [] \<rbrakk> \<Longrightarrow> Rangep R y"
by(induction rule: rtrancl_path.induct)(auto elim: rtrancl_path.cases)
lemma rtrancl_path_nth:
--- a/src/HOL/Library/Types_To_Sets.thy Wed Jan 04 14:26:08 2017 +0100
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,27 +0,0 @@
-(* Title: HOL/Library/Types_To_Sets.thy
- Author: Ondřej Kunčar, TU München
-*)
-
-section \<open>From Types to Sets\<close>
-
-text \<open>This theory extends Isabelle/HOL's logic by two new inference rules
- to allow translation of types to sets as described in
- O. Kunčar, A. Popescu: From Types to Sets by Local Type Definitions in Higher-Order Logic
- available at http://www21.in.tum.de/~kuncar/documents/kuncar-popescu-t2s2016-extended.pdf.\<close>
-
-theory Types_To_Sets
- imports Main
-begin
-
-subsection \<open>Rules\<close>
-
-text\<open>The following file implements the Local Typedef Rule (LT) and extends the logic by the rule.\<close>
-ML_file "Types_To_Sets/local_typedef.ML"
-
-text\<open>The following file implements the Unoverloading Rule (UO) and extends the logic by the rule.\<close>
-ML_file "Types_To_Sets/unoverloading.ML"
-
-text\<open>The following file implements a derived rule that internalizes type class annotations.\<close>
-ML_file "Types_To_Sets/internalize_sort.ML"
-
-end
\ No newline at end of file
--- a/src/HOL/Library/Types_To_Sets/internalize_sort.ML Wed Jan 04 14:26:08 2017 +0100
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,68 +0,0 @@
-(* Title: internalize_sort.ML
- Author: Ondřej Kunčar, TU München
-
- Implements a derived rule (by using Thm.unconstrainT) that internalizes
- type class annotations.
-*)
-
-
-(*
- \<phi>['a::{c_1, ..., c_n} / 'a]
----------------------------------------------------------------------
- class.c_1 ops_1 \<Longrightarrow> ... \<Longrightarrow> class.c_n ops_n \<Longrightarrow> \<phi>['a::type / 'a]
-
-where class.c is the locale predicate associated with type class c and
-ops are operations associated with type class c. For example:
-If c = semigroup_add, then ops = op-, op+, 0, uminus.
-If c = finite, then ops = TYPE('a::type).
-*)
-
-signature INTERNALIZE_SORT =
-sig
- val internalize_sort: ctyp -> thm -> typ * thm
- val internalize_sort_attr: typ -> attribute
-end;
-
-structure Internalize_Sort : INTERNALIZE_SORT =
-struct
-
-fun internalize_sort ctvar thm =
- let
- val thy = Thm.theory_of_thm thm;
- val tvar = Thm.typ_of ctvar;
-
- val ((_, assms, classes),_) = Logic.unconstrainT [] (Thm.prop_of thm);
-
- fun is_proper_class thy = can (Axclass.get_info thy); (* trick by FH *)
- fun reduce_to_non_proper_sort (TVar (name, sort)) =
- TVar (name, Sign.minimize_sort thy (filter_out (is_proper_class thy) (Sign.complete_sort thy sort)))
- | reduce_to_non_proper_sort _ = error "not supported";
-
- val data = (map fst classes) ~~ assms;
-
- val new_tvar = get_first (fn (tvar', ((ren_tvar, _), _)) => if tvar = tvar'
- then SOME (reduce_to_non_proper_sort ren_tvar) else NONE) data
- |> the_default tvar;
-
- fun localify class = Class.rules thy class |> snd |> Thm.transfer thy;
-
- fun discharge_of_class tvar class = Thm.of_class (Thm.global_ctyp_of thy tvar, class);
-
- val rules = map (fn (tvar', ((ren_tvar, class), _)) => if tvar = tvar'
- then (if Class.is_class thy class then localify class else discharge_of_class ren_tvar class)
- else discharge_of_class ren_tvar class) data;
- in
- (new_tvar, (Thm.unconstrainT (Thm.strip_shyps thm) OF rules) |> Drule.zero_var_indexes)
- end;
-
-val tvar = Args.context -- Args.typ >> (fn (_, v as TFree _) => Logic.varifyT_global v
- | (ctxt, t) => error ("Not a type variable: " ^ Syntax.string_of_typ ctxt t));
-
-fun internalize_sort_attr tvar =
- Thm.rule_attribute [] (fn context => fn thm =>
- (snd (internalize_sort (Thm.ctyp_of (Context.proof_of context) tvar) thm)));
-
-val _ = Context.>> (Context.map_theory (Attrib.setup @{binding internalize_sort}
- (tvar >> internalize_sort_attr) "internalize a sort"));
-
-end;
\ No newline at end of file
--- a/src/HOL/Library/Types_To_Sets/local_typedef.ML Wed Jan 04 14:26:08 2017 +0100
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,84 +0,0 @@
-(* Title: local_typedef.ML
- Author: Ondřej Kunčar, TU München
-
- Implements the Local Typedef Rule and extends the logic by the rule.
-*)
-
-signature LOCAL_TYPEDEF =
-sig
- val cancel_type_definition: thm -> thm
- val cancel_type_definition_attr: attribute
-end;
-
-structure Local_Typedef : LOCAL_TYPEDEF =
-struct
-
-(*
-Local Typedef Rule (LT)
-
- \<Gamma> \<turnstile> (\<exists>(Rep::\<beta> \<Rightarrow> \<tau>) Abs. type_definition Rep Abs A) \<Longrightarrow> \<phi>
-------------------------------------------------------------- [\<beta> not in \<phi>, \<Gamma>, A;
- \<Gamma> \<turnstile> A \<noteq> \<emptyset> \<Longrightarrow> \<phi> sort(\<beta>)=HOL.type]
-*)
-
-(** BEGINNING OF THE CRITICAL CODE **)
-
-fun dest_typedef (Const (@{const_name Ex}, _) $ Abs (_, _,
- (Const (@{const_name Ex}, _) $ Abs (_, Abs_type,
- (Const (@{const_name type_definition}, _)) $ Bound 1 $ Bound 0 $ set)))) =
- (Abs_type, set)
- | dest_typedef t = raise TERM ("dest_typedef", [t]);
-
-fun cancel_type_definition_cterm thm =
- let
- fun err msg = raise THM ("cancel_type_definition: " ^ msg, 0, [thm]);
-
- val thy = Thm.theory_of_thm thm;
- val prop = Thm.prop_of thm;
- val hyps = Thm.hyps_of thm;
-
- val _ = null (Thm.tpairs_of thm) orelse err "the theorem contains unsolved flex-flex pairs";
-
- val (typedef_assm, phi) = Logic.dest_implies prop
- handle TERM _ => err "the theorem is not an implication";
- val (abs_type, set) = typedef_assm |> HOLogic.dest_Trueprop |> dest_typedef
- handle TERM _ => err ("the assumption is not of the form "
- ^ quote "\<exists>Rep Abs. type_definition Rep Abs A");
-
- val (repT, absT) = Term.dest_funT abs_type;
- val _ = Term.is_TVar absT orelse err "the abstract type is not a schematic type variable";
- val (absT_name, sorts) = Term.dest_TVar absT;
-
- val typeS = Sign.defaultS thy;
- val _ = sorts = typeS orelse err ("the type " ^ quote (fst absT_name) ^ " is not of the sort "
- ^ quote (Syntax.string_of_sort_global thy typeS));
-
- fun contains_absT tm = member (op=) (Term.add_tvars tm []) (absT_name, sorts);
- fun err_contains_absT_in msg = err (msg ^ " contains the forbidden type "
- ^ quote (fst absT_name));
- val _ = not (contains_absT phi) orelse err_contains_absT_in "the conclusion";
- val _ = not (contains_absT set) orelse err_contains_absT_in "the set term";
- (* the following test is superfluous; the meta hypotheses cannot currently contain TVars *)
- val _ = not (exists contains_absT hyps) orelse err_contains_absT_in "one of the hypotheses";
-
- val not_empty_assm = HOLogic.mk_Trueprop
- (HOLogic.mk_not (HOLogic.mk_eq (set, HOLogic.mk_set repT [])));
- val prop = Logic.list_implies (hyps @ [not_empty_assm], phi);
- in
- Thm.global_cterm_of thy prop |> Thm.weaken_sorts (Thm.shyps_of thm)
- end;
-
-(** END OF THE CRITICAL CODE **)
-
-val (_, cancel_type_definition_oracle) = Context.>>> (Context.map_theory_result
- (Thm.add_oracle (@{binding cancel_type_definition}, cancel_type_definition_cterm)));
-
-fun cancel_type_definition thm =
- Drule.implies_elim_list (cancel_type_definition_oracle thm) (map Thm.assume (Thm.chyps_of thm));
-
-val cancel_type_definition_attr = Thm.rule_attribute [] (K cancel_type_definition);
-
-val _ = Context.>> (Context.map_theory (Attrib.setup @{binding cancel_type_definition}
- (Scan.succeed cancel_type_definition_attr) "cancel a local type definition"));
-
-end;
\ No newline at end of file
--- a/src/HOL/Library/Types_To_Sets/unoverloading.ML Wed Jan 04 14:26:08 2017 +0100
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,139 +0,0 @@
-(* Title: unoverloading.ML
- Author: Ondřej Kunčar, TU München
-
- Implements the Unoverloading Rule and extends the logic by the rule.
-*)
-
-signature UNOVERLOADING =
-sig
- val unoverload: cterm -> thm -> thm
- val unoverload_attr: cterm -> attribute
-end;
-
-structure Unoverloading : UNOVERLOADING =
-struct
-
-(*
-Unoverloading Rule (UO)
-
- \<turnstile> \<phi>[c::\<sigma> / x::\<sigma>]
----------------------------- [no type or constant or type class in \<phi>
- \<turnstile> \<And>x::\<sigma>. \<phi> depends on c::\<sigma>; c::\<sigma> is undefined]
-*)
-
-(* The following functions match_args, reduction and entries_of were
- cloned from defs.ML and theory.ML because the functions are not public.
- Notice that these functions already belong to the critical code.
-*)
-
-(* >= *)
-fun match_args (Ts, Us) =
- if Type.could_matches (Ts, Us) then
- Option.map Envir.subst_type
- (SOME (Type.raw_matches (Ts, Us) Vartab.empty) handle Type.TYPE_MATCH => NONE)
- else NONE;
-
-fun reduction defs (deps : Defs.entry list) : Defs.entry list option =
- let
- fun reduct Us (Ts, rhs) =
- (case match_args (Ts, Us) of
- NONE => NONE
- | SOME subst => SOME (map (apsnd (map subst)) rhs));
- fun reducts (d, Us) = get_first (reduct Us) (Defs.get_deps defs d);
-
- val reds = map (`reducts) deps;
- val deps' =
- if forall (is_none o #1) reds then NONE
- else SOME (fold_rev
- (fn (NONE, dp) => insert (op =) dp | (SOME dps, _) => fold (insert (op =)) dps) reds []);
- in deps' end;
-
-fun const_and_typ_entries_of thy tm =
- let
- val consts =
- fold_aterms (fn Const const => insert (op =) (Theory.const_dep thy const) | _ => I) tm [];
- val types =
- (fold_types o fold_subtypes) (fn Type t => insert (op =) (Theory.type_dep t) | _ => I) tm [];
- in
- consts @ types
- end;
-
-(* The actual implementation *)
-
-(** BEGINNING OF THE CRITICAL CODE **)
-
-fun fold_atyps_classes f =
- fold_atyps (fn T as TFree (_, S) => fold (pair T #> f) S
- | T as TVar (_, S) => fold (pair T #> f) S
- (* just to avoid a warning about incomplete patterns *)
- | _ => raise TERM ("fold_atyps_classes", []));
-
-fun class_entries_of thy tm =
- let
- val var_classes = (fold_types o fold_atyps_classes) (insert op=) tm [];
- in
- map (Logic.mk_of_class #> Term.head_of #> Term.dest_Const #> Theory.const_dep thy) var_classes
- end;
-
-fun unoverload_impl cconst thm =
- let
- fun err msg = raise THM ("unoverloading: " ^ msg, 0, [thm]);
-
- val _ = null (Thm.hyps_of thm) orelse err "the theorem has meta hypotheses";
- val _ = null (Thm.tpairs_of thm) orelse err "the theorem contains unresolved flex-flex pairs";
-
- val prop = Thm.prop_of thm;
- val prop_tfrees = rev (Term.add_tfree_names prop []);
- val _ = null prop_tfrees orelse err ("the theorem contains free type variables "
- ^ commas_quote prop_tfrees);
-
- val const = Thm.term_of cconst;
- val _ = Term.is_Const const orelse err "the parameter is is not a constant";
- val const_tfrees = rev (Term.add_tfree_names const []);
- val _ = null const_tfrees orelse err ("the constant contains free type variables "
- ^ commas_quote const_tfrees);
-
- val thy = Thm.theory_of_thm thm;
- val defs = Theory.defs_of thy;
-
- val const_entry = Theory.const_dep thy (Term.dest_Const const);
-
- val Uss = Defs.specifications_of defs (fst const_entry);
- val _ = forall (fn spec => is_none (match_args (#lhs spec, snd const_entry))) Uss
- orelse err "the constant instance has already a specification";
-
- val context = Defs.global_context thy;
- val prt_entry = Pretty.string_of o Defs.pretty_entry context;
-
- fun dep_err (c, Ts) (d, Us) =
- err (prt_entry (c, Ts) ^ " depends on " ^ prt_entry (d, Us));
- fun deps_of entry = perhaps_loop (reduction defs) [entry] |> these;
- fun not_depends_on_const prop_entry = forall (not_equal const_entry) (deps_of prop_entry)
- orelse dep_err prop_entry const_entry;
- val prop_entries = const_and_typ_entries_of thy prop @ class_entries_of thy prop;
- val _ = forall not_depends_on_const prop_entries;
- in
- Thm.global_cterm_of thy (Logic.all const prop) |> Thm.weaken_sorts (Thm.shyps_of thm)
- end;
-
-(** END OF THE CRITICAL CODE **)
-
-val (_, unoverload_oracle) = Context.>>> (Context.map_theory_result
- (Thm.add_oracle (@{binding unoverload},
- fn (const, thm) => unoverload_impl const thm)));
-
-fun unoverload const thm = unoverload_oracle (const, thm);
-
-fun unoverload_attr const =
- Thm.rule_attribute [] (fn _ => fn thm => (unoverload const thm));
-
-val const = Args.context -- Args.term >>
- (fn (ctxt, tm) =>
- if not (Term.is_Const tm)
- then error ("The term is not a constant: " ^ Syntax.string_of_term ctxt tm)
- else tm |> Logic.varify_types_global |> Thm.cterm_of ctxt);
-
-val _ = Context.>> (Context.map_theory (Attrib.setup @{binding unoverload}
- (const >> unoverload_attr) "unoverload an uninterpreted constant"));
-
-end;
\ No newline at end of file
--- a/src/HOL/Library/code_test.ML Wed Jan 04 14:26:08 2017 +0100
+++ b/src/HOL/Library/code_test.ML Wed Jan 04 14:26:19 2017 +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/Library/old_recdef.ML Wed Jan 04 14:26:08 2017 +0100
+++ b/src/HOL/Library/old_recdef.ML Wed Jan 04 14:26:19 2017 +0100
@@ -2779,15 +2779,15 @@
val recdef_wfN = "recdef_wf";
val recdef_modifiers =
- [Args.$$$ recdef_simpN -- Args.colon >> K (Method.modifier simp_add @{here}),
- Args.$$$ recdef_simpN -- Args.add -- Args.colon >> K (Method.modifier simp_add @{here}),
- Args.$$$ recdef_simpN -- Args.del -- Args.colon >> K (Method.modifier simp_del @{here}),
- Args.$$$ recdef_congN -- Args.colon >> K (Method.modifier cong_add @{here}),
- Args.$$$ recdef_congN -- Args.add -- Args.colon >> K (Method.modifier cong_add @{here}),
- Args.$$$ recdef_congN -- Args.del -- Args.colon >> K (Method.modifier cong_del @{here}),
- Args.$$$ recdef_wfN -- Args.colon >> K (Method.modifier wf_add @{here}),
- Args.$$$ recdef_wfN -- Args.add -- Args.colon >> K (Method.modifier wf_add @{here}),
- Args.$$$ recdef_wfN -- Args.del -- Args.colon >> K (Method.modifier wf_del @{here})] @
+ [Args.$$$ recdef_simpN -- Args.colon >> K (Method.modifier simp_add \<^here>),
+ Args.$$$ recdef_simpN -- Args.add -- Args.colon >> K (Method.modifier simp_add \<^here>),
+ Args.$$$ recdef_simpN -- Args.del -- Args.colon >> K (Method.modifier simp_del \<^here>),
+ Args.$$$ recdef_congN -- Args.colon >> K (Method.modifier cong_add \<^here>),
+ Args.$$$ recdef_congN -- Args.add -- Args.colon >> K (Method.modifier cong_add \<^here>),
+ Args.$$$ recdef_congN -- Args.del -- Args.colon >> K (Method.modifier cong_del \<^here>),
+ Args.$$$ recdef_wfN -- Args.colon >> K (Method.modifier wf_add \<^here>),
+ Args.$$$ recdef_wfN -- Args.add -- Args.colon >> K (Method.modifier wf_add \<^here>),
+ Args.$$$ recdef_wfN -- Args.del -- Args.colon >> K (Method.modifier wf_del \<^here>)] @
Clasimp.clasimp_modifiers;
--- a/src/HOL/Mirabelle/Tools/mirabelle.ML Wed Jan 04 14:26:08 2017 +0100
+++ b/src/HOL/Mirabelle/Tools/mirabelle.ML Wed Jan 04 14:26:19 2017 +0100
@@ -165,15 +165,17 @@
fun fold_body_thms f =
let
- fun app n (PBody {thms, ...}) = thms |> fold (fn (i, (name, prop, body)) =>
+ fun app n (PBody {thms, ...}) = thms |> fold (fn (i, thm_node) =>
fn (x, seen) =>
if Inttab.defined seen i then (x, seen)
else
let
- val body' = Future.join body
- val (x', seen') = app (n + (if name = "" then 0 else 1)) body'
+ val name = Proofterm.thm_node_name thm_node
+ val prop = Proofterm.thm_node_prop thm_node
+ val body = Future.join (Proofterm.thm_node_body thm_node)
+ val (x', seen') = app (n + (if name = "" then 0 else 1)) body
(x, Inttab.update (i, ()) seen)
- in (x' |> n = 0 ? f (name, prop, body'), seen') end)
+ in (x' |> n = 0 ? f (name, prop, body), seen') end)
in fn bodies => fn x => #1 (fold (app 0) bodies (x, Inttab.empty)) end
in
--- a/src/HOL/Nat.thy Wed Jan 04 14:26:08 2017 +0100
+++ b/src/HOL/Nat.thy Wed Jan 04 14:26:19 2017 +0100
@@ -2,9 +2,6 @@
Author: Tobias Nipkow
Author: Lawrence C Paulson
Author: Markus Wenzel
-
-Type "nat" is a linear order, and a datatype; arithmetic operators + -
-and * (for div and mod, see theory Divides).
*)
section \<open>Natural numbers\<close>
--- a/src/HOL/Nonstandard_Analysis/CLim.thy Wed Jan 04 14:26:08 2017 +0100
+++ b/src/HOL/Nonstandard_Analysis/CLim.thy Wed Jan 04 14:26:19 2017 +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 Wed Jan 04 14:26:08 2017 +0100
+++ b/src/HOL/Nonstandard_Analysis/CStar.thy Wed Jan 04 14:26:19 2017 +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 Wed Jan 04 14:26:08 2017 +0100
+++ b/src/HOL/Nonstandard_Analysis/Examples/NSPrimes.thy Wed Jan 04 14:26:19 2017 +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 Wed Jan 04 14:26:08 2017 +0100
+++ b/src/HOL/Nonstandard_Analysis/HDeriv.thy Wed Jan 04 14:26:19 2017 +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 Wed Jan 04 14:26:08 2017 +0100
+++ b/src/HOL/Nonstandard_Analysis/HLim.thy Wed Jan 04 14:26:19 2017 +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 Wed Jan 04 14:26:08 2017 +0100
+++ b/src/HOL/Nonstandard_Analysis/HLog.thy Wed Jan 04 14:26:19 2017 +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 Wed Jan 04 14:26:08 2017 +0100
+++ b/src/HOL/Nonstandard_Analysis/HSEQ.thy Wed Jan 04 14:26:19 2017 +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 Wed Jan 04 14:26:08 2017 +0100
+++ b/src/HOL/Nonstandard_Analysis/HSeries.thy Wed Jan 04 14:26:19 2017 +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 Wed Jan 04 14:26:08 2017 +0100
+++ b/src/HOL/Nonstandard_Analysis/StarDef.thy Wed Jan 04 14:26:19 2017 +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 Wed Jan 04 14:26:08 2017 +0100
+++ b/src/HOL/Number_Theory/Cong.thy Wed Jan 04 14:26:19 2017 +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 Wed Jan 04 14:26:08 2017 +0100
+++ b/src/HOL/Number_Theory/Euclidean_Algorithm.thy Wed Jan 04 14:26:19 2017 +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/Factorial_Ring.thy Wed Jan 04 14:26:08 2017 +0100
+++ b/src/HOL/Number_Theory/Factorial_Ring.thy Wed Jan 04 14:26:19 2017 +0100
@@ -451,7 +451,7 @@
lemma prime_dvd_multD: "prime p \<Longrightarrow> p dvd a * b \<Longrightarrow> p dvd a \<or> p dvd b"
by (intro prime_elem_dvd_multD) simp_all
-lemma prime_dvd_mult_iff [simp]: "prime p \<Longrightarrow> p dvd a * b \<longleftrightarrow> p dvd a \<or> p dvd b"
+lemma prime_dvd_mult_iff: "prime p \<Longrightarrow> p dvd a * b \<longleftrightarrow> p dvd a \<or> p dvd b"
by (auto dest: prime_dvd_multD)
lemma prime_dvd_power:
@@ -1640,7 +1640,6 @@
(insert nz False, auto simp: Gcd_factorial_eq_0_iff)
qed (simp_all add: Gcd_factorial_def)
-
lemma normalize_Lcm_factorial:
"normalize (Lcm_factorial A) = Lcm_factorial A"
proof (cases "subset_mset.bdd_above (prime_factorization ` A)")
--- a/src/HOL/Number_Theory/Gauss.thy Wed Jan 04 14:26:08 2017 +0100
+++ b/src/HOL/Number_Theory/Gauss.thy Wed Jan 04 14:26:19 2017 +0100
@@ -12,12 +12,12 @@
lemma cong_prime_prod_zero_nat:
fixes a::nat
shows "\<lbrakk>[a * b = 0] (mod p); prime p\<rbrakk> \<Longrightarrow> [a = 0] (mod p) | [b = 0] (mod p)"
- by (auto simp add: cong_altdef_nat)
+ by (auto simp add: cong_altdef_nat prime_dvd_mult_iff)
lemma cong_prime_prod_zero_int:
fixes a::int
shows "\<lbrakk>[a * b = 0] (mod p); prime p\<rbrakk> \<Longrightarrow> [a = 0] (mod p) | [b = 0] (mod p)"
- by (auto simp add: cong_altdef_int)
+ by (auto simp add: cong_altdef_int prime_dvd_mult_iff)
locale GAUSS =
--- a/src/HOL/Number_Theory/Pocklington.thy Wed Jan 04 14:26:08 2017 +0100
+++ b/src/HOL/Number_Theory/Pocklington.thy Wed Jan 04 14:26:19 2017 +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 Wed Jan 04 14:26:08 2017 +0100
+++ b/src/HOL/Number_Theory/Primes.thy Wed Jan 04 14:26:19 2017 +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 Wed Jan 04 14:26:08 2017 +0100
+++ b/src/HOL/Number_Theory/Quadratic_Reciprocity.thy Wed Jan 04 14:26:19 2017 +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 Wed Jan 04 14:26:08 2017 +0100
+++ b/src/HOL/Number_Theory/Residues.thy Wed Jan 04 14:26:19 2017 +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/Orderings.thy Wed Jan 04 14:26:08 2017 +0100
+++ b/src/HOL/Orderings.thy Wed Jan 04 14:26:19 2017 +0100
@@ -405,6 +405,10 @@
lemma not_le_imp_less: "\<not> y \<le> x \<Longrightarrow> x < y"
unfolding not_le .
+lemma linorder_less_wlog[case_names less refl sym]:
+ "\<lbrakk>\<And>a b. a < b \<Longrightarrow> P a b; \<And>a. P a a; \<And>a b. P b a \<Longrightarrow> P a b\<rbrakk> \<Longrightarrow> P a b"
+ using antisym_conv3 by blast
+
text \<open>Dual order\<close>
lemma dual_linorder:
--- a/src/HOL/Partial_Function.thy Wed Jan 04 14:26:08 2017 +0100
+++ b/src/HOL/Partial_Function.thy Wed Jan 04 14:26:19 2017 +0100
@@ -283,8 +283,8 @@
lemma flat_ord_antisym: "\<lbrakk> flat_ord a x y; flat_ord a y x \<rbrakk> \<Longrightarrow> x = y"
by(auto simp add: flat_ord_def)
-lemma antisymP_flat_ord: "antisymP (flat_ord a)"
-by(rule antisymI)(auto dest: flat_ord_antisym)
+lemma antisymp_flat_ord: "antisymp (flat_ord a)"
+by(rule antisympI)(auto dest: flat_ord_antisym)
interpretation tailrec:
partial_function_definitions "flat_ord undefined" "flat_lub undefined"
--- a/src/HOL/Power.thy Wed Jan 04 14:26:08 2017 +0100
+++ b/src/HOL/Power.thy Wed Jan 04 14:26:19 2017 +0100
@@ -582,10 +582,22 @@
context linordered_idom
begin
-lemma power_abs: "\<bar>a ^ n\<bar> = \<bar>a\<bar> ^ n"
- by (induct n) (auto simp add: abs_mult)
+lemma zero_le_power2 [simp]: "0 \<le> a\<^sup>2"
+ by (simp add: power2_eq_square)
+
+lemma zero_less_power2 [simp]: "0 < a\<^sup>2 \<longleftrightarrow> a \<noteq> 0"
+ by (force simp add: power2_eq_square zero_less_mult_iff linorder_neq_iff)
-lemma abs_power_minus [simp]: "\<bar>(-a) ^ n\<bar> = \<bar>a ^ n\<bar>"
+lemma power2_less_0 [simp]: "\<not> a\<^sup>2 < 0"
+ by (force simp add: power2_eq_square mult_less_0_iff)
+
+lemma power_abs: "\<bar>a ^ n\<bar> = \<bar>a\<bar> ^ n" -- \<open>FIXME simp?\<close>
+ by (induct n) (simp_all add: abs_mult)
+
+lemma power_sgn [simp]: "sgn (a ^ n) = sgn a ^ n"
+ by (induct n) (simp_all add: sgn_mult)
+
+lemma abs_power_minus [simp]: "\<bar>(- a) ^ n\<bar> = \<bar>a ^ n\<bar>"
by (simp add: power_abs)
lemma zero_less_power_abs_iff [simp]: "0 < \<bar>a\<bar> ^ n \<longleftrightarrow> a \<noteq> 0 \<or> n = 0"
@@ -600,15 +612,6 @@
lemma zero_le_power_abs [simp]: "0 \<le> \<bar>a\<bar> ^ n"
by (rule zero_le_power [OF abs_ge_zero])
-lemma zero_le_power2 [simp]: "0 \<le> a\<^sup>2"
- by (simp add: power2_eq_square)
-
-lemma zero_less_power2 [simp]: "0 < a\<^sup>2 \<longleftrightarrow> a \<noteq> 0"
- by (force simp add: power2_eq_square zero_less_mult_iff linorder_neq_iff)
-
-lemma power2_less_0 [simp]: "\<not> a\<^sup>2 < 0"
- by (force simp add: power2_eq_square mult_less_0_iff)
-
lemma power2_less_eq_zero_iff [simp]: "a\<^sup>2 \<le> 0 \<longleftrightarrow> a = 0"
by (simp add: le_less)
@@ -618,7 +621,7 @@
lemma power2_abs [simp]: "\<bar>a\<bar>\<^sup>2 = a\<^sup>2"
by (simp add: power2_eq_square)
-lemma odd_power_less_zero: "a < 0 \<Longrightarrow> a ^ Suc (2*n) < 0"
+lemma odd_power_less_zero: "a < 0 \<Longrightarrow> a ^ Suc (2 * n) < 0"
proof (induct n)
case 0
then show ?case by simp
@@ -630,11 +633,11 @@
by (simp del: power_Suc add: Suc mult_less_0_iff mult_neg_neg)
qed
-lemma odd_0_le_power_imp_0_le: "0 \<le> a ^ Suc (2*n) \<Longrightarrow> 0 \<le> a"
+lemma odd_0_le_power_imp_0_le: "0 \<le> a ^ Suc (2 * n) \<Longrightarrow> 0 \<le> a"
using odd_power_less_zero [of a n]
by (force simp add: linorder_not_less [symmetric])
-lemma zero_le_even_power'[simp]: "0 \<le> a ^ (2*n)"
+lemma zero_le_even_power'[simp]: "0 \<le> a ^ (2 * n)"
proof (induct n)
case 0
show ?case by simp
--- a/src/HOL/Predicate_Compile_Examples/Predicate_Compile_Tests.thy Wed Jan 04 14:26:08 2017 +0100
+++ b/src/HOL/Predicate_Compile_Examples/Predicate_Compile_Tests.thy Wed Jan 04 14:26:19 2017 +0100
@@ -125,6 +125,14 @@
code_pred predicate_where_argument_is_neg_condition_and_in_equation .
+text {* Another related example that required slight adjustment of the proof procedure *}
+
+inductive if_as_predicate :: "bool \<Rightarrow> 'a \<Rightarrow> 'a \<Rightarrow> 'a \<Rightarrow> bool"
+where
+ "condition \<Longrightarrow> if_as_predicate condition then_value else_value then_value"
+| "\<not> condition \<Longrightarrow> if_as_predicate condition then_value else_value else_value"
+
+code_pred [show_proof_trace] if_as_predicate .
inductive zerozero :: "nat * nat => bool"
where
--- a/src/HOL/Probability/Probability_Mass_Function.thy Wed Jan 04 14:26:08 2017 +0100
+++ b/src/HOL/Probability/Probability_Mass_Function.thy Wed Jan 04 14:26:19 2017 +0100
@@ -1573,34 +1573,29 @@
fixes p q :: "'a pmf"
assumes 1: "rel_pmf R p q"
assumes 2: "rel_pmf R q p"
- and refl: "reflp R" and trans: "transp R" and antisym: "antisymP R"
+ and refl: "reflp R" and trans: "transp R" and antisym: "antisymp R"
shows "p = q"
proof -
from 1 2 refl trans have "rel_pmf (inf R R\<inverse>\<inverse>) p q" by(rule rel_pmf_inf)
also have "inf R R\<inverse>\<inverse> = op ="
- using refl antisym by (auto intro!: ext simp add: reflpD dest: antisymD)
+ using refl antisym by (auto intro!: ext simp add: reflpD dest: antisympD)
finally show ?thesis unfolding pmf.rel_eq .
qed
lemma reflp_rel_pmf: "reflp R \<Longrightarrow> reflp (rel_pmf R)"
-by(blast intro: reflpI rel_pmf_reflI reflpD)
+ by (fact pmf.rel_reflp)
-lemma antisymP_rel_pmf:
- "\<lbrakk> reflp R; transp R; antisymP R \<rbrakk>
- \<Longrightarrow> antisymP (rel_pmf R)"
-by(rule antisymI)(blast intro: rel_pmf_antisym)
+lemma antisymp_rel_pmf:
+ "\<lbrakk> reflp R; transp R; antisymp R \<rbrakk>
+ \<Longrightarrow> antisymp (rel_pmf R)"
+by(rule antisympI)(blast intro: rel_pmf_antisym)
lemma transp_rel_pmf:
assumes "transp R"
shows "transp (rel_pmf R)"
-proof (rule transpI)
- fix x y z
- assume "rel_pmf R x y" and "rel_pmf R y z"
- hence "rel_pmf (R OO R) x z" by (simp add: pmf.rel_compp relcompp.relcompI)
- thus "rel_pmf R x z"
- using assms by (metis (no_types) pmf.rel_mono rev_predicate2D transp_relcompp_less_eq)
-qed
+ using assms by (fact pmf.rel_transp)
+
subsection \<open> Distributions \<close>
context
--- a/src/HOL/Probability/SPMF.thy Wed Jan 04 14:26:08 2017 +0100
+++ b/src/HOL/Probability/SPMF.thy Wed Jan 04 14:26:19 2017 +0100
@@ -95,8 +95,8 @@
lemma transp_ord_option: "transp ord \<Longrightarrow> transp (ord_option ord)"
unfolding transp_def by(blast intro: ord_option_trans)
-lemma antisymP_ord_option: "antisymP ord \<Longrightarrow> antisymP (ord_option ord)"
-by(auto intro!: antisymI elim!: ord_option.cases dest: antisymD)
+lemma antisymp_ord_option: "antisymp ord \<Longrightarrow> antisymp (ord_option ord)"
+by(auto intro!: antisympI elim!: ord_option.cases dest: antisympD)
lemma ord_option_chainD:
"Complete_Partial_Order.chain (ord_option ord) Y
@@ -1508,7 +1508,7 @@
fix x y
assume "?R x y" "?R y x"
thus "x = y"
- by(rule rel_pmf_antisym)(simp_all add: reflp_ord_option transp_ord_option antisymP_ord_option)
+ by(rule rel_pmf_antisym)(simp_all add: reflp_ord_option transp_ord_option antisymp_ord_option)
next
fix Y x
assume "Complete_Partial_Order.chain ?R Y" "x \<in> Y"
--- a/src/HOL/Proofs/ex/Proof_Terms.thy Wed Jan 04 14:26:08 2017 +0100
+++ b/src/HOL/Proofs/ex/Proof_Terms.thy Wed Jan 04 14:26:19 2017 +0100
@@ -30,7 +30,7 @@
(*all theorems used in the graph of nested proofs*)
val all_thms =
Proofterm.fold_body_thms
- (fn (name, _, _) => insert (op =) name) [body] [];
+ (fn {name, ...} => insert (op =) name) [body] [];
\<close>
text \<open>
--- a/src/HOL/Quickcheck_Exhaustive.thy Wed Jan 04 14:26:08 2017 +0100
+++ b/src/HOL/Quickcheck_Exhaustive.thy Wed Jan 04 14:26:19 2017 +0100
@@ -476,13 +476,18 @@
end
-(* FIXME instantiation char :: check_all
+instantiation char :: check_all
begin
-definition
- "check_all f = check_all (\<lambda>(x, t1). check_all (\<lambda>(y, t2).
- f (Char x y, \<lambda>_. Code_Evaluation.App
- (Code_Evaluation.App (Code_Evaluation.term_of Char) (t1 ())) (t2 ()))))"
+primrec check_all_char' ::
+ "(char \<times> (unit \<Rightarrow> term) \<Rightarrow> (bool \<times> term list) option) \<Rightarrow> char list \<Rightarrow> (bool \<times> term list) option"
+ where "check_all_char' f [] = None"
+ | "check_all_char' f (c # cs) = f (c, \<lambda>_. Code_Evaluation.term_of c)
+ orelse check_all_char' f cs"
+
+definition check_all_char ::
+ "(char \<times> (unit \<Rightarrow> term) \<Rightarrow> (bool \<times> term list) option) \<Rightarrow> (bool \<times> term list) option"
+ where "check_all f = check_all_char' f Enum.enum"
definition enum_term_of_char :: "char itself \<Rightarrow> unit \<Rightarrow> term list"
where
@@ -490,7 +495,7 @@
instance ..
-end *)
+end
instantiation option :: (check_all) check_all
begin
--- a/src/HOL/ROOT Wed Jan 04 14:26:08 2017 +0100
+++ b/src/HOL/ROOT Wed Jan 04 14:26:19 2017 +0100
@@ -18,10 +18,9 @@
description {*
HOL-Main with explicit proof terms.
*}
- options [document = false, quick_and_dirty = false]
+ options [document = false, quick_and_dirty = false, parallel_proofs = 0]
theories Proofs (*sequential change of global flag!*)
- theories List
- theories [checkpoint] "~~/src/HOL/Library/Old_Datatype"
+ theories "~~/src/HOL/Library/Old_Datatype"
files
"Tools/Quickcheck/Narrowing_Engine.hs"
"Tools/Quickcheck/PNF_Narrowing_Engine.hs"
@@ -32,19 +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
- Types_To_Sets
- (*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
@@ -56,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 +
@@ -240,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
@@ -396,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
@@ -1018,6 +1013,17 @@
theories [condition = ISABELLE_SWIPL, quick_and_dirty]
Reg_Exp_Example
+session "HOL-Types_To_Sets" in Types_To_Sets = HOL +
+ description {*
+ Experimental extension of Higher-Order Logic to allow translation of types to sets.
+ *}
+ options [document = false]
+ theories
+ Types_To_Sets
+ "Examples/Prerequisites"
+ "Examples/Finite"
+ "Examples/T2_Spaces"
+
session HOLCF (main timing) in HOLCF = HOL +
description {*
Author: Franz Regensburger
--- a/src/HOL/Rat.thy Wed Jan 04 14:26:08 2017 +0100
+++ b/src/HOL/Rat.thy Wed Jan 04 14:26:19 2017 +0100
@@ -824,9 +824,15 @@
lemma Rats_of_int [simp]: "of_int z \<in> \<rat>"
by (subst of_rat_of_int_eq [symmetric]) (rule Rats_of_rat)
+lemma Ints_subset_Rats: "\<int> \<subseteq> \<rat>"
+ using Ints_cases Rats_of_int by blast
+
lemma Rats_of_nat [simp]: "of_nat n \<in> \<rat>"
by (subst of_rat_of_nat_eq [symmetric]) (rule Rats_of_rat)
+lemma Nats_subset_Rats: "\<nat> \<subseteq> \<rat>"
+ using Ints_subset_Rats Nats_subset_Ints by blast
+
lemma Rats_number_of [simp]: "numeral w \<in> \<rat>"
by (subst of_rat_numeral_eq [symmetric]) (rule Rats_of_rat)
--- a/src/HOL/Relation.thy Wed Jan 04 14:26:08 2017 +0100
+++ b/src/HOL/Relation.thy Wed Jan 04 14:26:19 2017 +0100
@@ -321,26 +321,51 @@
definition antisym :: "'a rel \<Rightarrow> bool"
where "antisym r \<longleftrightarrow> (\<forall>x y. (x, y) \<in> r \<longrightarrow> (y, x) \<in> r \<longrightarrow> x = y)"
-abbreviation antisymP :: "('a \<Rightarrow> 'a \<Rightarrow> bool) \<Rightarrow> bool"
- where "antisymP r \<equiv> antisym {(x, y). r x y}" (* FIXME proper logical operation *)
+definition antisymp :: "('a \<Rightarrow> 'a \<Rightarrow> bool) \<Rightarrow> bool"
+ where "antisymp r \<longleftrightarrow> (\<forall>x y. r x y \<longrightarrow> r y x \<longrightarrow> x = y)"
-lemma antisymI [intro?]: "(\<And>x y. (x, y) \<in> r \<Longrightarrow> (y, x) \<in> r \<Longrightarrow> x = y) \<Longrightarrow> antisym r"
+lemma antisymp_antisym_eq [pred_set_conv]: "antisymp (\<lambda>x y. (x, y) \<in> r) \<longleftrightarrow> antisym r"
+ by (simp add: antisym_def antisymp_def)
+
+lemma antisymI [intro?]:
+ "(\<And>x y. (x, y) \<in> r \<Longrightarrow> (y, x) \<in> r \<Longrightarrow> x = y) \<Longrightarrow> antisym r"
unfolding antisym_def by iprover
-lemma antisymD [dest?]: "antisym r \<Longrightarrow> (a, b) \<in> r \<Longrightarrow> (b, a) \<in> r \<Longrightarrow> a = b"
+lemma antisympI [intro?]:
+ "(\<And>x y. r x y \<Longrightarrow> r y x \<Longrightarrow> x = y) \<Longrightarrow> antisymp r"
+ by (fact antisymI [to_pred])
+
+lemma antisymD [dest?]:
+ "antisym r \<Longrightarrow> (a, b) \<in> r \<Longrightarrow> (b, a) \<in> r \<Longrightarrow> a = b"
unfolding antisym_def by iprover
-lemma antisym_subset: "r \<subseteq> s \<Longrightarrow> antisym s \<Longrightarrow> antisym r"
- unfolding antisym_def by blast
+lemma antisympD [dest?]:
+ "antisymp r \<Longrightarrow> r a b \<Longrightarrow> r b a \<Longrightarrow> a = b"
+ by (fact antisymD [to_pred])
-lemma antisym_empty [simp]: "antisym {}"
+lemma antisym_subset:
+ "r \<subseteq> s \<Longrightarrow> antisym s \<Longrightarrow> antisym r"
unfolding antisym_def by blast
-lemma antisymP_equality [simp]: "antisymP op ="
- by (auto intro: antisymI)
+lemma antisymp_less_eq:
+ "r \<le> s \<Longrightarrow> antisymp s \<Longrightarrow> antisymp r"
+ by (fact antisym_subset [to_pred])
+
+lemma antisym_empty [simp]:
+ "antisym {}"
+ unfolding antisym_def by blast
-lemma antisym_singleton [simp]: "antisym {x}"
-by (blast intro: antisymI)
+lemma antisym_bot [simp]:
+ "antisymp \<bottom>"
+ by (fact antisym_empty [to_pred])
+
+lemma antisymp_equality [simp]:
+ "antisymp HOL.eq"
+ by (auto intro: antisympI)
+
+lemma antisym_singleton [simp]:
+ "antisym {x}"
+ by (blast intro: antisymI)
subsubsection \<open>Transitivity\<close>
@@ -389,8 +414,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)
@@ -436,21 +462,45 @@
definition single_valued :: "('a \<times> 'b) set \<Rightarrow> bool"
where "single_valued r \<longleftrightarrow> (\<forall>x y. (x, y) \<in> r \<longrightarrow> (\<forall>z. (x, z) \<in> r \<longrightarrow> y = z))"
-abbreviation single_valuedP :: "('a \<Rightarrow> 'b \<Rightarrow> bool) \<Rightarrow> bool"
- where "single_valuedP r \<equiv> single_valued {(x, y). r x y}" (* FIXME proper logical operation *)
+definition single_valuedp :: "('a \<Rightarrow> 'b \<Rightarrow> bool) \<Rightarrow> bool"
+ where "single_valuedp r \<longleftrightarrow> (\<forall>x y. r x y \<longrightarrow> (\<forall>z. r x z \<longrightarrow> y = z))"
+
+lemma single_valuedp_single_valued_eq [pred_set_conv]:
+ "single_valuedp (\<lambda>x y. (x, y) \<in> r) \<longleftrightarrow> single_valued r"
+ by (simp add: single_valued_def single_valuedp_def)
-lemma single_valuedI: "\<forall>x y. (x, y) \<in> r \<longrightarrow> (\<forall>z. (x, z) \<in> r \<longrightarrow> y = z) \<Longrightarrow> single_valued r"
- unfolding single_valued_def .
+lemma single_valuedI:
+ "(\<And>x y. (x, y) \<in> r \<Longrightarrow> (\<And>z. (x, z) \<in> r \<Longrightarrow> y = z)) \<Longrightarrow> single_valued r"
+ unfolding single_valued_def by blast
-lemma single_valuedD: "single_valued r \<Longrightarrow> (x, y) \<in> r \<Longrightarrow> (x, z) \<in> r \<Longrightarrow> y = z"
+lemma single_valuedpI:
+ "(\<And>x y. r x y \<Longrightarrow> (\<And>z. r x z \<Longrightarrow> y = z)) \<Longrightarrow> single_valuedp r"
+ by (fact single_valuedI [to_pred])
+
+lemma single_valuedD:
+ "single_valued r \<Longrightarrow> (x, y) \<in> r \<Longrightarrow> (x, z) \<in> r \<Longrightarrow> y = z"
by (simp add: single_valued_def)
-lemma single_valued_empty[simp]: "single_valued {}"
+lemma single_valuedpD:
+ "single_valuedp r \<Longrightarrow> r x y \<Longrightarrow> r x z \<Longrightarrow> y = z"
+ by (fact single_valuedD [to_pred])
+
+lemma single_valued_empty [simp]:
+ "single_valued {}"
by (simp add: single_valued_def)
-lemma single_valued_subset: "r \<subseteq> s \<Longrightarrow> single_valued s \<Longrightarrow> single_valued r"
+lemma single_valuedp_bot [simp]:
+ "single_valuedp \<bottom>"
+ by (fact single_valued_empty [to_pred])
+
+lemma single_valued_subset:
+ "r \<subseteq> s \<Longrightarrow> single_valued s \<Longrightarrow> single_valued r"
unfolding single_valued_def by blast
+lemma single_valuedp_less_eq:
+ "r \<le> s \<Longrightarrow> single_valuedp s \<Longrightarrow> single_valuedp r"
+ by (fact single_valued_subset [to_pred])
+
subsection \<open>Relation operations\<close>
@@ -620,13 +670,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
@@ -1146,12 +1198,4 @@
(auto simp: comp_fun_commute.fold_insert comp_fun_commute_relcomp_fold insert_relcomp_fold
cong: if_cong)
-text \<open>Misc\<close>
-
-abbreviation (input) transP :: "('a \<Rightarrow> 'a \<Rightarrow> bool) \<Rightarrow> bool"
- where "transP r \<equiv> trans {(x, y). r x y}" (* FIXME drop *)
-
-abbreviation (input) "RangeP \<equiv> Rangep"
-abbreviation (input) "DomainP \<equiv> Domainp"
-
end
--- a/src/HOL/Rings.thy Wed Jan 04 14:26:08 2017 +0100
+++ b/src/HOL/Rings.thy Wed Jan 04 14:26:19 2017 +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
@@ -2021,6 +2126,49 @@
lemma abs_sgn_eq: "\<bar>sgn a\<bar> = (if a = 0 then 0 else 1)"
by (simp add: sgn_if)
+lemma sgn_mult_self_eq [simp]:
+ "sgn a * sgn a = of_bool (a \<noteq> 0)"
+ by (cases "a > 0") simp_all
+
+lemma abs_mult_self_eq [simp]:
+ "\<bar>a\<bar> * \<bar>a\<bar> = a * a"
+ by (cases "a > 0") simp_all
+
+lemma same_sgn_sgn_add:
+ "sgn (a + b) = sgn a" if "sgn b = sgn a"
+proof (cases a 0 rule: linorder_cases)
+ case equal
+ with that show ?thesis
+ by simp
+next
+ case less
+ with that have "b < 0"
+ by (simp add: sgn_1_neg)
+ with \<open>a < 0\<close> have "a + b < 0"
+ by (rule add_neg_neg)
+ with \<open>a < 0\<close> show ?thesis
+ by simp
+next
+ case greater
+ with that have "b > 0"
+ by (simp add: sgn_1_pos)
+ with \<open>a > 0\<close> have "a + b > 0"
+ by (rule add_pos_pos)
+ with \<open>a > 0\<close> show ?thesis
+ by simp
+qed
+
+lemma same_sgn_abs_add:
+ "\<bar>a + b\<bar> = \<bar>a\<bar> + \<bar>b\<bar>" if "sgn b = sgn a"
+proof -
+ have "a + b = sgn a * \<bar>a\<bar> + sgn b * \<bar>b\<bar>"
+ by (simp add: sgn_mult_abs)
+ also have "\<dots> = sgn a * (\<bar>a\<bar> + \<bar>b\<bar>)"
+ using that by (simp add: algebra_simps)
+ finally show ?thesis
+ by (auto simp add: abs_mult)
+qed
+
lemma abs_dvd_iff [simp]: "\<bar>m\<bar> dvd k \<longleftrightarrow> m dvd k"
by (simp add: abs_if)
--- a/src/HOL/SPARK/Examples/RIPEMD-160/Round.thy Wed Jan 04 14:26:08 2017 +0100
+++ b/src/HOL/SPARK/Examples/RIPEMD-160/Round.thy Wed Jan 04 14:26:19 2017 +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 Wed Jan 04 14:26:08 2017 +0100
+++ b/src/HOL/SPARK/Manual/Proc1.thy Wed Jan 04 14:26:19 2017 +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 Wed Jan 04 14:26:08 2017 +0100
+++ b/src/HOL/SPARK/Manual/Proc2.thy Wed Jan 04 14:26:19 2017 +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/SPARK/Tools/spark_vcs.ML Wed Jan 04 14:26:08 2017 +0100
+++ b/src/HOL/SPARK/Tools/spark_vcs.ML Wed Jan 04 14:26:19 2017 +0100
@@ -293,7 +293,7 @@
| SOME _ => error ("Cannot associate a type with " ^ s ^
"\nsince it is no record or enumeration type");
-fun check_enum [] [] = NONE
+fun check_enum [] [] = NONE
| check_enum els [] = SOME ("has no element(s) " ^ commas els)
| check_enum [] cs = SOME ("has extra element(s) " ^
commas (map (Long_Name.base_name o fst) cs))
@@ -305,7 +305,7 @@
fun invert_map [] = I
| invert_map cmap =
map (apfst (the o AList.lookup (op =) (map swap cmap)));
-
+
fun add_type_def prfx (s, Basic_Type ty) (ids, thy) =
(check_no_assoc thy prfx s;
(ids,
@@ -677,7 +677,7 @@
"+", "-", "*", "/", "div", "mod", "**"]);
fun complex_expr (Number _) = false
- | complex_expr (Ident _) = false
+ | complex_expr (Ident _) = false
| complex_expr (Funct (s, es)) =
not (Symtab.defined builtin s) orelse exists complex_expr es
| complex_expr (Quantifier (_, _, _, e)) = complex_expr e
@@ -959,7 +959,7 @@
| SOME {vcs, path, ...} =>
let
val (proved, unproved) = partition_vcs vcs;
- val _ = Thm.join_proofs (maps (#2 o snd) proved);
+ val _ = List.app Thm.consolidate (maps (#2 o snd) proved);
val (proved', proved'') =
List.partition (fn (_, (_, thms, _, _)) =>
exists (#oracle o Thm.peek_status) thms) proved;
@@ -1117,7 +1117,7 @@
[(term_of_rule thy' prfx types pfuns ids rl, [])]))
other_rules),
Element.Notes ("", [((Binding.name "defns", []), map (rpair [] o single o snd) defs')])]
-
+
in
set_env ctxt defs' types funs ids vcs' path prfx thy'
end;
--- a/src/HOL/Statespace/state_fun.ML Wed Jan 04 14:26:08 2017 +0100
+++ b/src/HOL/Statespace/state_fun.ML Wed Jan 04 14:26:19 2017 +0100
@@ -77,7 +77,7 @@
fun string_eq_simp_tac ctxt =
simp_tac (put_simpset HOL_basic_ss ctxt
addsimps @{thms list.inject list.distinct Char_eq_Char_iff
- cut_eq_simps simp_thms}
+ cong_exp_iff_simps simp_thms}
addsimprocs [lazy_conj_simproc]
|> Simplifier.add_cong @{thm block_conj_cong});
--- a/src/HOL/String.thy Wed Jan 04 14:26:08 2017 +0100
+++ b/src/HOL/String.thy Wed Jan 04 14:26:19 2017 +0100
@@ -114,7 +114,7 @@
"nat_of_char (Char k) = numeral k mod 256"
by (simp add: Char_def)
-lemma Char_eq_Char_iff [simp]:
+lemma Char_eq_Char_iff:
"Char k = Char l \<longleftrightarrow> numeral k mod (256 :: nat) = numeral l mod 256" (is "?P \<longleftrightarrow> ?Q")
proof -
have "?P \<longleftrightarrow> nat_of_char (Char k) = nat_of_char (Char l)"
@@ -124,14 +124,25 @@
finally show ?thesis .
qed
-lemma zero_eq_Char_iff [simp]:
+lemma zero_eq_Char_iff:
"0 = Char k \<longleftrightarrow> numeral k mod (256 :: nat) = 0"
by (auto simp add: zero_char_def Char_def)
-lemma Char_eq_zero_iff [simp]:
+lemma Char_eq_zero_iff:
"Char k = 0 \<longleftrightarrow> numeral k mod (256 :: nat) = 0"
by (auto simp add: zero_char_def Char_def)
+simproc_setup char_eq
+ ("Char m = Char n" | "Char m = 0" | "0 = Char n") = \<open>
+ let
+ val ss = put_simpset HOL_ss @{context}
+ |> fold Simplifier.add_simp @{thms Char_eq_Char_iff zero_eq_Char_iff Char_eq_zero_iff cong_exp_iff_simps}
+ |> simpset_of
+ in
+ fn _ => fn ctxt => SOME o Simplifier.rewrite (put_simpset ss ctxt)
+ end
+\<close>
+
definition integer_of_char :: "char \<Rightarrow> integer"
where
"integer_of_char = integer_of_nat \<circ> nat_of_char"
--- a/src/HOL/TPTP/ATP_Problem_Import.thy Wed Jan 04 14:26:08 2017 +0100
+++ b/src/HOL/TPTP/ATP_Problem_Import.thy Wed Jan 04 14:26:19 2017 +0100
@@ -22,6 +22,7 @@
(*
ML {* ATP_Problem_Import.isabelle_tptp_file @{theory} 50 "$TPTP/Problems/PUZ/PUZ107^5.p" *}
+ML {* ATP_Problem_Import.isabelle_hot_tptp_file @{theory} 20 "$TPTP/Problems/SYO/SYO304^5.p" *}
*)
end
--- a/src/HOL/TPTP/CASC/ReadMe Wed Jan 04 14:26:08 2017 +0100
+++ b/src/HOL/TPTP/CASC/ReadMe Wed Jan 04 14:26:19 2017 +0100
@@ -1,4 +1,4 @@
-Notes from Geoff:
+Notes from Geoff Sutcliffe:
I added a few lines to the top of bin/isabelle ...
@@ -48,10 +48,11 @@
STAREXEC_WALLCLOCK_LIMIT=300 ./bin/starexec_run_isabelle_hot $TPTP/Problems/CSR/CSR150^3.p
STAREXEC_WALLCLOCK_LIMIT=300 ./bin/starexec_run_isabelle_hot $TPTP/Problems/SYO/SYO304^5.p
- The first problem is unprovable; the second one is proved by Satallax.
+ The first problem is unprovable; the second one is proved by Satallax (after
+ some delay).
- All the tools accept CNF, FOF, TFF0, or THF0 problems and output SZS
- statuses of the form
+ All the tools accept CNF, FOF, TFF0, TFF1, THF0, or THF1 problems and output
+ SZS statuses of the form
% SZS status XXX
@@ -60,43 +61,32 @@
{Unknown, TimedOut, Unsatisfiable, Theorem, Satisfiable, CounterSatisfiable}
Nitpick also output a model within "% SZS begin" and "% SZS end" tags, in
- its idiosyncratic syntax.
-
- In 2011, there were some problems with Java (needed for Nitpick), because it
- required so much memory at startup. I doubt there will be any problems this
- year, because Isabelle now includes its own version of Java, but the solution
- back then was to replace
-
- exec isabelle java
-
- in the last line of the "contrib/kodkodi-1.5.2/bin/kodkodi" script with
-
- /usr/lib64/jvm/java-1.5.0-gcj-4.5-1.5.0.0/jre/bin/java
-
- See the emails we exchanged on 18 July 2011, with the subject "No problem on
- my Linux 64-bit".
+ its idiosyncratic syntax. For TFF0 and THF0, phantom type arguments are not
+ supported, and type quantifiers are only allowed at the outermost position
+ in a formula, as "forall".
Enjoy!
Notes to myself:
- I downloaded the official Isabelle2015 Linux package from
+ I downloaded the official Isabelle2016-1 Linux package from
- http://isabelle.in.tum.de/dist/Isabelle2015_linux.tar.gz
+ http://isabelle.in.tum.de/dist/Isabelle2016-1_linux.tar.gz
- on "macbroy21" and renamed the directory "Isabelle2015-CASC". I modified
+ on "macbroy21" and renamed the directory "Isabelle2016-1-CASC". I modified
- src/HOL/TPTP/atp_problem_import.ML
+ src/HOL/TPTP
to include changes backported from the development version of Isabelle. I
- then built a "HOL-TPTP" image:
+ also modified "bin/isabelle" as suggested by Geoff above. I then built a
+ "HOL-TPTP" image:
./bin/isabelle build -b HOL-TPTP
- I copied the heaps over to "./heaps":
+ I moved the heaps over to "./heaps":
- mv ~/.isabelle/Isabelle2015/heaps .
+ mv ~/.isabelle/Isabelle2016-1/heaps .
I created some wrapper scripts in "./bin":
@@ -117,7 +107,7 @@
http://page.mi.fu-berlin.de/cbenzmueller/leo/leo2_v1.6.2.tgz
I did "make opt". I copied "bin/leo.opt" to
- "~/Isabelle2015-CASC/contrib/leo".
+ "~/Isabelle2016-1-CASC/contrib/leo".
I added this line to "etc/settings":
@@ -141,19 +131,19 @@
SATALLAX_HOME=$ISABELLE_HOME/contrib
- Vampire (2.6):
+ Vampire 4.0 (commit 2fedff6)
- I copied the file "vampire", which I probably got from the 2013 CASC
- archive and moved it to "~/Isabelle2013-CASC/contrib/vampire".
+ I copied the file "vampire", which I got from Giles Reger on 23 September
+ 2015.
I added these lines to "etc/settings":
VAMPIRE_HOME=$ISABELLE_HOME/contrib
- VAMPIRE_VERSION=3.0
+ VAMPIRE_VERSION=4.0
Z3 TPTP (4.3.2.0 postrelease):
- I cloned out the git repository:
+ For Isabelle2015, I cloned out the git repository:
git clone https://git01.codeplex.com/z3
@@ -173,9 +163,11 @@
"/tmp/T.thy" with the following content:
theory T imports Main begin
+
lemma "a = b ==> [b] = [a]"
- sledgehammer [cvc4 e leo2 satallax spass vampire z3 z3_tptp] ()
- oops
+ sledgehammer [cvc4 e leo2 satallax spass vampire z3 (*z3_tptp*)] ()
+ oops
+
end
Then I ran
@@ -196,4 +188,4 @@
Jasmin Blanchette
- 10 June 2015
+ 15 December 2016
--- a/src/HOL/TPTP/CASC/SysDesc_Isabelle.html Wed Jan 04 14:26:08 2017 +0100
+++ b/src/HOL/TPTP/CASC/SysDesc_Isabelle.html Wed Jan 04 14:26:19 2017 +0100
@@ -1,5 +1,5 @@
<HR><!------------------------------------------------------------------------>
-<H2>Isabelle/HOL 2015</H2>
+<H2>Isabelle/HOL 2016-1</H2>
Jasmin C. Blanchette<sup>1</sup>, Lawrence C. Paulson<sup>2</sup>,
Tobias Nipkow<sup>1</sup>, Makarius Wenzel<sup>3</sup> <BR>
<sup>1</sup>Technische Universität München, Germany <BR>
@@ -8,9 +8,9 @@
<H3>Architecture</H3>
-Isabelle/HOL 2015 [<A HREF="#References">NPW13</A>] is the higher-order
-logic incarnation of the generic proof assistant
-<A HREF="http://www.cl.cam.ac.uk/research/hvg/Isabelle/">Isabelle2015</A>.
+Isabelle/HOL 2016-1 [<A HREF="#References">NPW13</A>] is the higher-order
+logic incarnation of the generic proof assistant
+<A HREF="http://www.cl.cam.ac.uk/research/hvg/Isabelle/">Isabelle2016-1</A>.
Isabelle/HOL provides several automatic proof tactics, notably an equational
reasoner [<A HREF="#References">Nip89</A>], a classical reasoner [<A
HREF="#References">Pau94</A>], and a tableau prover [<A
--- a/src/HOL/TPTP/CASC/SysDesc_Nitpick.html Wed Jan 04 14:26:08 2017 +0100
+++ b/src/HOL/TPTP/CASC/SysDesc_Nitpick.html Wed Jan 04 14:26:19 2017 +0100
@@ -1,11 +1,11 @@
<HR><!------------------------------------------------------------------------>
-<H2>Nitpick 2015</H2>
+<H2>Nitpick 2016-1</H2>
Jasmin C. Blanchette<BR>
Technische Universität München, Germany <BR>
<H3>Architecture</H3>
-Nitpick [<A HREF="#References">BN10</A>] is an open source counterexample
+Nitpick [<A HREF="#References">BN10</A>] is an open source counterexample
generator for Isabelle/HOL [<A HREF="#References">NPW13</A>]. It builds on
Kodkod [<A HREF="#References">TJ07</A>], a highly optimized first-order
relational model finder based on SAT. The name Nitpick is appropriated from a
--- a/src/HOL/TPTP/CASC/SysDesc_Refute.html Wed Jan 04 14:26:08 2017 +0100
+++ b/src/HOL/TPTP/CASC/SysDesc_Refute.html Wed Jan 04 14:26:19 2017 +0100
@@ -1,12 +1,12 @@
<HR><!------------------------------------------------------------------------>
-<H2>Refute 2015</H2>
+<H2>Refute 2016-1</H2>
Jasmin C. Blanchette<sup>1</sup>, Tjark Weber<sup>2</sup><BR>
<sup>1</sup>Technische Universität München, Germany <BR>
<sup>2</sup>Uppsala Universitet, Sweden <BR>
<H3>Architecture</H3>
-Refute [<A HREF="#References">Web08</A>] is an open source counterexample
+Refute [<A HREF="#References">Web08</A>] is an open source counterexample
generator for Isabelle/HOL [<A HREF="#References">NPW13</A>] based on a
SAT solver, and Nitpick's [<A HREF="#References">BN10</A>] precursor.
--- a/src/HOL/TPTP/TPTP_Parser/tptp_interpret.ML Wed Jan 04 14:26:08 2017 +0100
+++ b/src/HOL/TPTP/TPTP_Parser/tptp_interpret.ML Wed Jan 04 14:26:19 2017 +0100
@@ -215,7 +215,7 @@
fun declare_constant config const_name ty thy =
if const_exists config thy const_name then
raise MISINTERPRET_TERM
- ("Const already declared", Term_Func (Uninterpreted const_name, []))
+ ("Const already declared", Term_FuncG (Uninterpreted const_name, [], []))
else
Theory.specify_const ((mk_binding config const_name, ty), NoSyn) thy
@@ -224,10 +224,10 @@
(*Types in TFF/THF are encoded as formulas. These functions translate them to type form.*)
-fun termtype_to_type (Term_Func (TypeSymbol typ, [])) =
+fun termtype_to_type (Term_FuncG (TypeSymbol typ, [], [])) =
Defined_type typ
- | termtype_to_type (Term_Func (Uninterpreted str, tms)) =
- Atom_type (str, map termtype_to_type tms)
+ | termtype_to_type (Term_FuncG (Uninterpreted str, tys, tms)) =
+ Atom_type (str, tys @ map termtype_to_type tms)
| termtype_to_type (Term_Var str) = Var_type str
(*FIXME possibly incomplete hack*)
@@ -249,6 +249,9 @@
| fmlatype_to_type (Quant (Dep_Prod, _, fmla)) = fmlatype_to_type fmla
| fmlatype_to_type (Pred (Interpreted_ExtraLogic Apply, [tm])) =
termtype_to_type tm
+ | fmlatype_to_type (Fmla (Interpreted_ExtraLogic Apply, [tm1, tm2])) =
+ (case fmlatype_to_type tm1 of
+ Atom_type (str, tys) => Atom_type (str, tys @ [fmlatype_to_type tm2]))
fun tfree_name_of_var_type str = "'" ^ Name.desymbolize (SOME false) str
fun tfree_of_var_type str = TFree (tfree_name_of_var_type str, @{sort type})
@@ -323,7 +326,7 @@
else
raise MISINTERPRET_TERM
("Could not find the interpretation of this constant in the \
- \mapping to Isabelle/HOL", Term_Func (Uninterpreted str, [])))
+ \mapping to Isabelle/HOL", Term_FuncG (Uninterpreted str, [], [])))
(*Eta-expands n-ary function.
"str" is the name of an Isabelle/HOL constant*)
@@ -448,7 +451,7 @@
interpret_type config thy type_map (Defined_type Type_Bool)
else ind_type
in case tptp_atom_value of
- Atom_App (Term_Func (symb, tptp_ts)) =>
+ Atom_App (Term_FuncG (symb, tptp_tys, tptp_ts)) =>
let
val thy' = fold (fn t =>
type_atoms_to_thy config false type_map (Atom_App t)) tptp_ts thy
@@ -456,7 +459,8 @@
case symb of
Uninterpreted const_name =>
perhaps (try (snd o declare_constant config const_name
- (mk_fun_type (replicate (length tptp_ts) ind_type) value_type I))) thy'
+ (mk_fun_type (replicate (length tptp_tys + length tptp_ts) ind_type) value_type I)))
+ thy'
| _ => thy'
end
| Atom_App _ => thy
@@ -499,6 +503,25 @@
| _ => false
end
+fun strip_applies_term (Term_FuncG (Interpreted_ExtraLogic Apply, [], [tm1, tm2])) =
+ strip_applies_term tm1 ||> (fn tms => tms @ [tm2])
+ | strip_applies_term tm = (tm, [])
+
+fun termify_apply_fmla thy config (Fmla (Interpreted_ExtraLogic Apply, [fmla1, fmla2])) =
+ (case termify_apply_fmla thy config fmla1 of
+ SOME (Term_FuncG (symb, tys, tms)) =>
+ let val typ_arity = type_arity_of_symbol thy config symb in
+ (case (null tms andalso length tys < typ_arity, try fmlatype_to_type fmla2) of
+ (true, SOME ty) => SOME (Term_FuncG (symb, tys @ [ty], []))
+ | _ =>
+ (case termify_apply_fmla thy config fmla2 of
+ SOME tm2 => SOME (Term_FuncG (symb, tys, tms @ [tm2]))
+ | NONE => NONE))
+ end
+ | _ => NONE)
+ | termify_apply_fmla _ _ (Atom (THF_Atom_term tm)) = SOME tm
+ | termify_apply_fmla _ _ _ = NONE
+
(*
Information needed to be carried around:
- constant mapping: maps constant names to Isabelle terms with fully-qualified
@@ -508,29 +531,33 @@
after each call of interpret_term since variables' cannot be bound across
terms.
*)
-fun interpret_term formula_level config language const_map var_types type_map
- tptp_t thy =
+fun interpret_term formula_level config language const_map var_types type_map tptp_t thy =
case tptp_t of
- Term_Func (symb, tptp_ts) =>
+ Term_FuncG (symb, tptp_tys, tptp_ts) =>
let
val thy' =
type_atoms_to_thy config formula_level type_map (Atom_App tptp_t) thy
in
case symb of
Interpreted_ExtraLogic Apply =>
+ (case strip_applies_term tptp_t of
+ (Term_FuncG (symb, tptp_tys, tptp_ts), extra_tptp_ts) =>
+ interpret_term formula_level config language const_map var_types type_map
+ (Term_FuncG (symb, tptp_tys, tptp_ts @ extra_tptp_ts)) thy
+ | _ =>
(*apply the head of the argument list to the tail*)
(mapply'
(fold_map (interpret_term false config language const_map
var_types type_map) (tl tptp_ts) thy')
(interpret_term formula_level config language const_map
- var_types type_map (hd tptp_ts)))
- | _ =>
+ var_types type_map (hd tptp_ts))))
+ | _ =>
let
val typ_arity = type_arity_of_symbol thy' config symb
- val (tptp_type_args, tptp_term_args) = chop typ_arity tptp_ts
+ val (tptp_type_args, tptp_term_args) = chop (typ_arity - length tptp_tys) tptp_ts
val tyargs =
- map (interpret_type config thy' type_map o termtype_to_type)
- tptp_type_args
+ map (interpret_type config thy' type_map)
+ (tptp_tys @ map termtype_to_type tptp_type_args)
in
(*apply symb to tptp_ts*)
if is_prod_typed thy' config symb then
@@ -604,38 +631,40 @@
and interpret_formula config language const_map var_types type_map tptp_fmla thy =
case tptp_fmla of
- Pred app =>
+ Pred (symb, ts) =>
interpret_term true config language const_map
- var_types type_map (Term_Func app) thy
+ var_types type_map (Term_FuncG (symb, [], ts)) thy
| Fmla (symbol, tptp_formulas) =>
(case symbol of
Interpreted_ExtraLogic Apply =>
+ (case termify_apply_fmla thy config tptp_fmla of
+ SOME tptp_t =>
+ interpret_term true config language const_map var_types type_map tptp_t thy
+ | NONE =>
mapply'
(fold_map (interpret_formula config language const_map
var_types type_map) (tl tptp_formulas) thy)
(interpret_formula config language const_map
- var_types type_map (hd tptp_formulas))
+ var_types type_map (hd tptp_formulas)))
| Uninterpreted const_name =>
let
val (args, thy') = (fold_map (interpret_formula config language
- const_map var_types type_map) tptp_formulas thy)
+ const_map var_types type_map) tptp_formulas thy)
val thy' =
type_atoms_to_thy config true type_map
- (Atom_Arity (const_name, length tptp_formulas)) thy'
+ (Atom_Arity (const_name, length tptp_formulas)) thy'
in
(if is_prod_typed thy' config symbol then
mtimes thy' args (interpret_symbol config const_map symbol [] thy')
else
- mapply args (interpret_symbol config const_map symbol [] thy'),
+ mapply args (interpret_symbol config const_map symbol [] thy'),
thy')
end
| _ =>
let
val (args, thy') =
- fold_map
- (interpret_formula config language
- const_map var_types type_map)
- tptp_formulas thy
+ fold_map (interpret_formula config language const_map var_types type_map)
+ tptp_formulas thy
in
(if is_prod_typed thy' config symbol then
mtimes thy' args (interpret_symbol config const_map symbol [] thy')
@@ -737,19 +766,20 @@
| _ => ([], tptp_type)
in
fun typeof_typing (THF_typing (_, tptp_type)) = extract_type tptp_type
- fun typeof_tff_typing (Atom (TFF_Typed_Atom (_, SOME tptp_type))) =
- extract_type tptp_type
+ | typeof_typing (Atom (TFF_Typed_Atom (_, SOME tptp_type))) = extract_type tptp_type
end
-fun nameof_typing
- (THF_typing
- ((Atom (THF_Atom_term (Term_Func (Uninterpreted str, [])))), _)) = str
-fun nameof_tff_typing (Atom (TFF_Typed_Atom (Uninterpreted str, _))) = str
+fun nameof_typing (THF_typing
+ ((Atom (THF_Atom_term (Term_FuncG (Uninterpreted str, [], [])))), _)) = str
+ | nameof_typing (Atom (TFF_Typed_Atom (Uninterpreted str, _))) = str
fun strip_prod_type (Prod_type (ty1, ty2)) = ty1 :: strip_prod_type ty2
| strip_prod_type ty = [ty]
-fun dest_fn_type (Fn_type (ty1, ty2)) = (strip_prod_type ty1, ty2)
+fun dest_fn_type (Fn_type (ty1, ty2)) =
+ let val (tys2, ty3) = dest_fn_type ty2 in
+ (strip_prod_type ty1 @ tys2, ty3)
+ end
| dest_fn_type ty = ([], ty)
fun resolve_include_path path_prefixes path_suffix =
@@ -758,11 +788,16 @@
SOME prefix => Path.append prefix path_suffix
| NONE => error ("Cannot find include file " ^ Path.print path_suffix)
-(* Ideally, to be in sync with TFF1, we should perform full type skolemization here.
- But the problems originating from HOL systems are restricted to top-level
- universal quantification for types. *)
+fun is_type_type (Fmla_type (Atom (THF_Atom_term (Term_FuncG (TypeSymbol Type_Type, [], []))))) =
+ true
+ | is_type_type (Defined_type Type_Type) = true
+ | is_type_type _ = false;
+
+(* Ideally, to be in sync with TFF1/THF1, we should perform full type
+ skolemization here. But the problems originating from HOL systems are
+ restricted to top-level universal quantification for types. *)
fun remove_leading_type_quantifiers (Quant (Forall, varlist, tptp_fmla)) =
- (case filter_out (curry (op =) (SOME (Defined_type Type_Type)) o snd) varlist of
+ (case filter_out (fn (_, SOME ty) => is_type_type ty | _ => false) varlist of
[] => remove_leading_type_quantifiers tptp_fmla
| varlist' => Quant (Forall, varlist', tptp_fmla))
| remove_leading_type_quantifiers tptp_fmla = tptp_fmla
@@ -788,12 +823,8 @@
Role_Type =>
let
val ((tptp_type_vars, tptp_ty), name) =
- if lang = THF then
- (typeof_typing tptp_formula (*assuming tptp_formula is a typing*),
- nameof_typing tptp_formula (*and that the LHS is a (atom) name*))
- else
- (typeof_tff_typing tptp_formula,
- nameof_tff_typing tptp_formula)
+ (typeof_typing tptp_formula (*assuming tptp_formula is a typing*),
+ nameof_typing tptp_formula (*and that the LHS is a (atom) name*))
in
case dest_fn_type tptp_ty of
(tptp_binders, Defined_type Type_Type) => (*add new type*)
@@ -865,7 +896,7 @@
(*gather interpreted term, and the map of types occurring in that term*)
val (t, thy') =
interpret_formula config lang
- const_map [] type_map (remove_leading_type_quantifiers tptp_formula) thy
+ const_map [] type_map (remove_leading_type_quantifiers tptp_formula) thy
|>> HOLogic.mk_Trueprop
(*type_maps grow monotonically, so return the newest value (type_mapped);
there's no need to unify it with the type_map parameter.*)
--- a/src/HOL/TPTP/TPTP_Parser/tptp_syntax.ML Wed Jan 04 14:26:08 2017 +0100
+++ b/src/HOL/TPTP/TPTP_Parser/tptp_syntax.ML Wed Jan 04 14:26:19 2017 +0100
@@ -81,7 +81,7 @@
| General_List of general_term list
and tptp_term =
- Term_Func of symbol * tptp_term list
+ Term_FuncG of symbol * tptp_type list (*special hack for TPTP_Interpret*) * tptp_term list
| Term_Var of string
| Term_Conditional of tptp_formula * tptp_term * tptp_term
| Term_Num of number_kind * string
@@ -118,6 +118,8 @@
| Fmla_type of tptp_formula
| Subtype of symbol * symbol (*only THF*)
+ val Term_Func: symbol * tptp_term list -> tptp_term (*for Yacc parser*)
+
type general_list = general_term list
type parent_details = general_list
type useful_info = general_term list
@@ -230,7 +232,7 @@
| General_List of general_term list
and tptp_term =
- Term_Func of symbol * tptp_term list
+ Term_FuncG of symbol * tptp_type list (*special hack for TPTP_Interpret*) * tptp_term list
| Term_Var of string
| Term_Conditional of tptp_formula * tptp_term * tptp_term
| Term_Num of number_kind * string
@@ -267,6 +269,8 @@
| Fmla_type of tptp_formula
| Subtype of symbol * symbol
+fun Term_Func (symb, ts) = Term_FuncG (symb, [], ts)
+
type general_list = general_term list
type parent_details = general_list
type useful_info = general_term list
@@ -405,9 +409,10 @@
fun string_of_tptp_term x =
case x of
- Term_Func (symbol, tptp_term_list) =>
+ Term_FuncG (symbol, tptp_type_list, tptp_term_list) =>
"(" ^ string_of_symbol symbol ^ " " ^
- space_implode " " (map string_of_tptp_term tptp_term_list) ^ ")"
+ space_implode " " (map string_of_tptp_type tptp_type_list
+ @ map string_of_tptp_term tptp_term_list) ^ ")"
| Term_Var str => str
| Term_Conditional (tptp_formula, tptp_term1, tptp_term2) => "" (*FIXME*)
| Term_Num (_, str) => str
@@ -531,22 +536,23 @@
(*infix symbols, including \subset, \cup, \cap*)
fun latex_of_tptp_term x =
case x of
- Term_Func (Interpreted_Logic Equals, [tptp_t1, tptp_t2]) =>
+ Term_FuncG (Interpreted_Logic Equals, [], [tptp_t1, tptp_t2]) =>
"(" ^ latex_of_tptp_term tptp_t1 ^ " = " ^ latex_of_tptp_term tptp_t2 ^ ")"
- | Term_Func (Interpreted_Logic NEquals, [tptp_t1, tptp_t2]) =>
+ | Term_FuncG (Interpreted_Logic NEquals, [], [tptp_t1, tptp_t2]) =>
"(" ^ latex_of_tptp_term tptp_t1 ^ " \\\\neq " ^ latex_of_tptp_term tptp_t2 ^ ")"
- | Term_Func (Interpreted_Logic Or, [tptp_t1, tptp_t2]) =>
+ | Term_FuncG (Interpreted_Logic Or, [], [tptp_t1, tptp_t2]) =>
"(" ^ latex_of_tptp_term tptp_t1 ^ " \\\\vee " ^ latex_of_tptp_term tptp_t2 ^ ")"
- | Term_Func (Interpreted_Logic And, [tptp_t1, tptp_t2]) =>
+ | Term_FuncG (Interpreted_Logic And, [], [tptp_t1, tptp_t2]) =>
"(" ^ latex_of_tptp_term tptp_t1 ^ " \\\\wedge " ^ latex_of_tptp_term tptp_t2 ^ ")"
- | Term_Func (Interpreted_Logic Iff, [tptp_t1, tptp_t2]) =>
+ | Term_FuncG (Interpreted_Logic Iff, [], [tptp_t1, tptp_t2]) =>
"(" ^ latex_of_tptp_term tptp_t1 ^ " \\\\longleftrightarrow " ^ latex_of_tptp_term tptp_t2 ^ ")"
- | Term_Func (Interpreted_Logic If, [tptp_t1, tptp_t2]) =>
+ | Term_FuncG (Interpreted_Logic If, [], [tptp_t1, tptp_t2]) =>
"(" ^ latex_of_tptp_term tptp_t1 ^ " \\\\longrightarrow " ^ latex_of_tptp_term tptp_t2 ^ ")"
- | Term_Func (symbol, tptp_term_list) =>
+ | Term_FuncG (symbol, tptp_type_list, tptp_term_list) =>
(*"(" ^*) latex_of_symbol symbol ^ "\\\\, " ^
- space_implode "\\\\, " (map latex_of_tptp_term tptp_term_list) (*^ ")"*)
+ space_implode "\\\\, " (map latex_of_tptp_type tptp_type_list
+ @ map latex_of_tptp_term tptp_term_list) (*^ ")"*)
| Term_Var str => "\\\\mathrm{" ^ str ^ "}"
| Term_Conditional (tptp_formula, tptp_term1, tptp_term2) => "" (*FIXME*)
| Term_Num (_, str) => str
@@ -647,10 +653,10 @@
latex_of_symbol symbol ^
space_implode "\\\\, " (map latex_of_tptp_term tptp_term_list)
- | latex_of_tptp_formula (Fmla (Interpreted_ExtraLogic Apply, [Fmla (Interpreted_ExtraLogic Apply, [Atom (THF_Atom_term (Term_Func (Uninterpreted "union", []))), tptp_f1]), tptp_f2])) =
+ | latex_of_tptp_formula (Fmla (Interpreted_ExtraLogic Apply, [Fmla (Interpreted_ExtraLogic Apply, [Atom (THF_Atom_term (Term_FuncG (Uninterpreted "union", [], []))), tptp_f1]), tptp_f2])) =
"(" ^ latex_of_tptp_formula tptp_f1 ^ " \\\\cup " ^ latex_of_tptp_formula tptp_f2 ^ ")"
- | latex_of_tptp_formula (Fmla (Interpreted_ExtraLogic Apply, [Fmla (Interpreted_ExtraLogic Apply, [Atom (THF_Atom_term (Term_Func (Uninterpreted "subset", []))), tptp_f1]), tptp_f2])) =
+ | latex_of_tptp_formula (Fmla (Interpreted_ExtraLogic Apply, [Fmla (Interpreted_ExtraLogic Apply, [Atom (THF_Atom_term (Term_FuncG (Uninterpreted "subset", [], []))), tptp_f1]), tptp_f2])) =
"(" ^ latex_of_tptp_formula tptp_f1 ^ " \\\\subset " ^ latex_of_tptp_formula tptp_f2 ^ ")"
| latex_of_tptp_formula (Fmla (Interpreted_Logic Equals, [tptp_f1, tptp_f2])) =
--- a/src/HOL/TPTP/TPTP_Parser/tptp_to_dot.ML Wed Jan 04 14:26:08 2017 +0100
+++ b/src/HOL/TPTP/TPTP_Parser/tptp_to_dot.ML Wed Jan 04 14:26:19 2017 +0100
@@ -96,7 +96,7 @@
which we interpret to be the last line of the proof.*)
fun is_last_line CNF (Pred (Interpreted_Logic False, [])) = true
| is_last_line THF (Atom (THF_Atom_term
- (Term_Func (Interpreted_Logic False, [])))) = true
+ (Term_Func (Interpreted_Logic False, [], [])))) = true
| is_last_line _ _ = false
fun tptp_dot_node with_label reverse_arrows
--- a/src/HOL/TPTP/atp_problem_import.ML Wed Jan 04 14:26:08 2017 +0100
+++ b/src/HOL/TPTP/atp_problem_import.ML Wed Jan 04 14:26:19 2017 +0100
@@ -11,7 +11,7 @@
'a list * ('a list * 'a list) * Proof.context
val nitpick_tptp_file : theory -> int -> string -> unit
val refute_tptp_file : theory -> int -> string -> unit
- val can_tac : Proof.context -> tactic -> term -> bool
+ val can_tac : Proof.context -> (Proof.context -> tactic) -> term -> bool
val SOLVE_TIMEOUT : int -> string -> tactic -> tactic
val atp_tac : Proof.context -> int -> (string * string) list -> int -> term list -> string ->
int -> tactic
@@ -131,14 +131,15 @@
(** Sledgehammer and Isabelle (combination of provers) **)
-fun can_tac ctxt tactic conj = can (Goal.prove ctxt [] [] conj) (K tactic)
+fun can_tac ctxt tactic conj =
+ can (Goal.prove_internal ctxt [] (Thm.cterm_of ctxt conj)) (fn [] => tactic ctxt)
fun SOLVE_TIMEOUT seconds name tac st =
let
val _ = writeln ("running " ^ name ^ " for " ^ string_of_int seconds ^ " s")
val result =
Timeout.apply (Time.fromSeconds seconds) (fn () => SINGLE (SOLVE tac) st) ()
- handle Timeout.TIMEOUT _ => NONE | ERROR _ => NONE
+ handle Timeout.TIMEOUT _ => NONE | ERROR _ => NONE
in
(case result of
NONE => (writeln ("FAILURE: " ^ name); Seq.empty)
@@ -274,7 +275,7 @@
val conj = make_conj ([], []) conjs
val assms = op @ assms
in
- can_tac ctxt (sledgehammer_tac true ctxt timeout assms 1) conj
+ can_tac ctxt (fn ctxt => sledgehammer_tac true ctxt timeout assms 1) conj
|> print_szs_of_success conjs
end
@@ -287,9 +288,9 @@
val (last_hope_atp, last_hope_completeness) =
if demo then (ATP_Proof.satallaxN, 0) else (ATP_Proof.vampireN, 2)
in
- (can_tac ctxt (auto_etc_tac ctxt (timeout div 2) assms 1) full_conj orelse
- can_tac ctxt (sledgehammer_tac demo ctxt (timeout div 2) assms 1) conj orelse
- can_tac ctxt (SOLVE_TIMEOUT timeout (last_hope_atp ^ "(*)")
+ (can_tac ctxt (fn ctxt => auto_etc_tac ctxt (timeout div 2) assms 1) full_conj orelse
+ can_tac ctxt (fn ctxt => sledgehammer_tac demo ctxt (timeout div 2) assms 1) conj orelse
+ can_tac ctxt (fn ctxt => SOLVE_TIMEOUT timeout (last_hope_atp ^ "(*)")
(atp_tac ctxt last_hope_completeness [] timeout assms last_hope_atp 1)) full_conj)
|> print_szs_of_success conjs
end
--- a/src/HOL/TPTP/lib/Tools/tptp_isabelle Wed Jan 04 14:26:08 2017 +0100
+++ b/src/HOL/TPTP/lib/Tools/tptp_isabelle Wed Jan 04 14:26:19 2017 +0100
@@ -24,12 +24,12 @@
TIMEOUT=$1
shift
-isabelle build -b HOL-TPTP
+isabelle build -b HOL-TPTP 2>&1 | grep --line-buffered -v "elapsed time$"
for FILE in "$@"
do
echo "theory $SCRATCH imports \"$TPTP_HOME/ATP_Problem_Import\" begin \
ML {* ATP_Problem_Import.isabelle_tptp_file @{theory} ($TIMEOUT) \"$FILE\" *} end" \
> /tmp/$SCRATCH.thy
- isabelle process -e "use_thy \"/tmp/$SCRATCH\"; exit 1;" -l HOL-TPTP | grep --line-buffered -v "^###\|^PROOF FAILED for depth\|^Failure node\|inferences so far. Searching to depth\|^val \|^Loading theory\|^Warning-The type of\|^ monotype.$"
+ isabelle process -e "use_thy \"/tmp/$SCRATCH\"; exit 1;" -l HOL-TPTP | grep --line-buffered -v "^###\|^PROOF FAILED for depth\|^Failure node\|inferences so far. Searching to depth\|^val \|^Loading theory\|^poly.*warning: The type of\|^ monotype.$"
done
--- a/src/HOL/TPTP/lib/Tools/tptp_isabelle_hot Wed Jan 04 14:26:08 2017 +0100
+++ b/src/HOL/TPTP/lib/Tools/tptp_isabelle_hot Wed Jan 04 14:26:19 2017 +0100
@@ -24,12 +24,12 @@
TIMEOUT=$1
shift
-isabelle build -b HOL-TPTP
+isabelle build -b HOL-TPTP 2>&1 | grep --line-buffered -v "elapsed time$"
for FILE in "$@"
do
echo "theory $SCRATCH imports \"$TPTP_HOME/ATP_Problem_Import\" begin \
ML {* ATP_Problem_Import.isabelle_hot_tptp_file @{theory} ($TIMEOUT) \"$FILE\" *} end" \
> /tmp/$SCRATCH.thy
- isabelle process -e "use_thy \"/tmp/$SCRATCH\"; exit 1;" -l HOL-TPTP | grep --line-buffered -v "^###\|^PROOF FAILED for depth\|^Failure node\|inferences so far. Searching to depth\|^val \|^Loading theory\|^Warning-The type of\|^ monotype.$"
+ isabelle process -e "use_thy \"/tmp/$SCRATCH\"; exit 1;" -l HOL-TPTP | grep --line-buffered -v "^###\|^PROOF FAILED for depth\|^Failure node\|inferences so far. Searching to depth\|^val \|^Loading theory\|^poly.*warning: The type of\|^ monotype.$"
done
--- a/src/HOL/TPTP/lib/Tools/tptp_nitpick Wed Jan 04 14:26:08 2017 +0100
+++ b/src/HOL/TPTP/lib/Tools/tptp_nitpick Wed Jan 04 14:26:19 2017 +0100
@@ -24,12 +24,12 @@
TIMEOUT=$1
shift
-isabelle build -b HOL-TPTP
+isabelle build -b HOL-TPTP 2>&1 | grep --line-buffered -v "elapsed time$"
for FILE in "$@"
do
echo "theory $SCRATCH imports \"$TPTP_HOME/ATP_Problem_Import\" begin \
ML {* ATP_Problem_Import.nitpick_tptp_file @{theory} ($TIMEOUT) \"$FILE\" *} end" \
> /tmp/$SCRATCH.thy
- isabelle process -e "use_thy \"/tmp/$SCRATCH\"; exit 1;" -l HOL-TPTP | grep --line-buffered -v "^###\|^PROOF FAILED for depth\|^Failure node\|inferences so far. Searching to depth\|^val \|^Loading theory\|^Warning-The type of\|^ monotype.$"
+ isabelle process -e "use_thy \"/tmp/$SCRATCH\"; exit 1;" -l HOL-TPTP | grep --line-buffered -v "^###\|^PROOF FAILED for depth\|^Failure node\|inferences so far. Searching to depth\|^val \|^Loading theory\|^poly.*warning: The type of\|^ monotype.$"
done
--- a/src/HOL/TPTP/lib/Tools/tptp_refute Wed Jan 04 14:26:08 2017 +0100
+++ b/src/HOL/TPTP/lib/Tools/tptp_refute Wed Jan 04 14:26:19 2017 +0100
@@ -23,12 +23,12 @@
TIMEOUT=$1
shift
-isabelle build -b HOL-TPTP
+isabelle build -b HOL-TPTP 2>&1 | grep --line-buffered -v "elapsed time$"
for FILE in "$@"
do
echo "theory $SCRATCH imports \"$TPTP_HOME/ATP_Problem_Import\" begin \
ML {* ATP_Problem_Import.refute_tptp_file @{theory} ($TIMEOUT) \"$FILE\" *} end" \
> /tmp/$SCRATCH.thy
- isabelle process -e "use_thy \"/tmp/$SCRATCH\"; exit 1;" -l HOL-TPTP | grep --line-buffered -v "^###\|^PROOF FAILED for depth\|^Failure node\|inferences so far. Searching to depth\|^val \|^Loading theory\|^Warning-The type of\|^ monotype.$"
+ isabelle process -e "use_thy \"/tmp/$SCRATCH\"; exit 1;" -l HOL-TPTP | grep --line-buffered -v "^###\|^PROOF FAILED for depth\|^Failure node\|inferences so far. Searching to depth\|^val \|^Loading theory\|^poly.*warning: The type of\|^ monotype.$"
done
--- a/src/HOL/TPTP/lib/Tools/tptp_sledgehammer Wed Jan 04 14:26:08 2017 +0100
+++ b/src/HOL/TPTP/lib/Tools/tptp_sledgehammer Wed Jan 04 14:26:19 2017 +0100
@@ -24,12 +24,12 @@
TIMEOUT=$1
shift
-isabelle build -b HOL-TPTP
+isabelle build -b HOL-TPTP 2>&1 | grep --line-buffered -v "elapsed time$"
for FILE in "$@"
do
echo "theory $SCRATCH imports \"$TPTP_HOME/ATP_Problem_Import\" begin \
ML {* ATP_Problem_Import.sledgehammer_tptp_file @{theory} ($TIMEOUT) \"$FILE\" *} end" \
> /tmp/$SCRATCH.thy
- isabelle process -e "use_thy \"/tmp/$SCRATCH\"; exit 1;" -l HOL-TPTP | grep --line-buffered -v "^###\|^PROOF FAILED for depth\|^Failure node\|inferences so far. Searching to depth\|^val \|^Loading theory\|^Warning-The type of\|^ monotype.$"
+ isabelle process -e "use_thy \"/tmp/$SCRATCH\"; exit 1;" -l HOL-TPTP | grep --line-buffered -v "^###\|^PROOF FAILED for depth\|^Failure node\|inferences so far. Searching to depth\|^val \|^Loading theory\|^poly.*warning: The type of\|^ monotype.$"
done
--- a/src/HOL/TPTP/lib/Tools/tptp_translate Wed Jan 04 14:26:08 2017 +0100
+++ b/src/HOL/TPTP/lib/Tools/tptp_translate Wed Jan 04 14:26:19 2017 +0100
@@ -23,9 +23,9 @@
args=("$@")
-isabelle build -b HOL-TPTP
+isabelle build -b HOL-TPTP 2>&1 | grep --line-buffered -v "elapsed time$"
echo "theory $SCRATCH imports \"$TPTP_HOME/ATP_Problem_Import\" begin \
ML {* ATP_Problem_Import.translate_tptp_file @{theory} \"${args[0]}\" \"${args[1]}\" *} end" \
> /tmp/$SCRATCH.thy
-isabelle process -e "use_thy \"/tmp/$SCRATCH\"; exit 1;" -l HOL-TPTP | grep --line-buffered -v "^###\|^PROOF FAILED for depth\|^Failure node\|inferences so far. Searching to depth\|^val \|^Loading theory\|^Warning-The type of\|^ monotype.$"
+isabelle process -e "use_thy \"/tmp/$SCRATCH\"; exit 1;" -l HOL-TPTP | grep --line-buffered -v "^###\|^PROOF FAILED for depth\|^Failure node\|inferences so far. Searching to depth\|^val \|^Loading theory\|^poly.*warning: The type of\|^ monotype.$"
--- a/src/HOL/Tools/BNF/bnf_def.ML Wed Jan 04 14:26:08 2017 +0100
+++ b/src/HOL/Tools/BNF/bnf_def.ML Wed Jan 04 14:26:19 2017 +0100
@@ -122,9 +122,9 @@
val mk_pred: typ list -> term -> term
val mk_rel: int -> typ list -> typ list -> term -> term
val mk_set: typ list -> term -> term
- val build_map: Proof.context -> typ list -> (typ * typ -> term) -> typ * typ -> term
- val build_rel: (string * (int * term)) list -> Proof.context -> typ list -> (typ * typ -> term) ->
- typ * typ -> term
+ val build_map: Proof.context -> typ list -> typ list -> (typ * typ -> term) -> typ * typ -> term
+ val build_rel: (string * (int * term)) list -> Proof.context -> typ list -> typ list ->
+ (typ * typ -> term) -> typ * typ -> term
val build_set: Proof.context -> typ -> typ -> term
val flatten_type_args_of_bnf: bnf -> 'a -> 'a list -> 'a list
val map_flattened_map_args: Proof.context -> string -> (term list -> 'a list) -> term list ->
@@ -744,12 +744,13 @@
Term.subst_atomic_types (Ts0 @ Us0 ~~ Ts @ Us) t
end;
-fun build_map_or_rel mk const of_bnf dest pre_cst_table ctxt simple_Ts build_simple =
+fun build_map_or_rel mk const of_bnf dest pre_cst_table ctxt simple_Ts simple_Us build_simple =
let
fun build (TU as (T, U)) =
- if exists (curry (op =) T) simple_Ts then
+ if exists (curry (op =) T) simple_Ts orelse exists (curry (op =) U) simple_Us then
build_simple TU
- else if T = U andalso not (exists_subtype_in simple_Ts T) then
+ else if T = U andalso not (exists_subtype_in simple_Ts T) andalso
+ not (exists_subtype_in simple_Us U) then
const T
else
(case TU of
--- a/src/HOL/Tools/BNF/bnf_fp_def_sugar.ML Wed Jan 04 14:26:08 2017 +0100
+++ b/src/HOL/Tools/BNF/bnf_fp_def_sugar.ML Wed Jan 04 14:26:19 2017 +0100
@@ -100,6 +100,7 @@
'a list
val mk_ctor: typ list -> term -> term
val mk_dtor: typ list -> term -> term
+ val mk_bnf_sets: BNF_Def.bnf -> string * term list
val liveness_of_fp_bnf: int -> BNF_Def.bnf -> bool list
val nesting_bnfs: Proof.context -> typ list list list -> typ list -> BNF_Def.bnf list
@@ -162,6 +163,14 @@
* (((term list list * term list list * term list list list list * term list list list list)
* term list list list) * typ list) -> (string -> binding) -> 'b list -> typ list ->
term list -> term -> local_theory -> (term * thm) * local_theory
+ val mk_induct_raw_prem: (term -> term) -> Proof.context -> typ list list ->
+ (string * term list) list -> term -> term -> typ list -> typ list ->
+ term list * ((term * (term * term)) list * (int * term)) list * term
+ val finish_induct_prem: Proof.context -> int -> term list ->
+ term list * ((term * (term * term)) list * (int * term)) list * term -> term
+ val mk_coinduct_prem: Proof.context -> typ list list -> typ list list -> term list -> term ->
+ term -> term -> int -> term list -> term list list -> term list -> term list list ->
+ typ list list -> term
val mk_induct_attrs: term list list -> Token.src list
val mk_coinduct_attrs: typ list -> term list list -> term list list -> int list list ->
Token.src list * Token.src list
@@ -170,18 +179,18 @@
BNF_Def.bnf list -> BNF_Def.bnf list -> typ list -> typ list -> typ list ->
typ list list list -> thm list -> thm list -> thm list -> term list list -> thm list list ->
term list -> thm list -> Proof.context -> lfp_sugar_thms
- val derive_coinduct_thms_for_types: bool -> (term -> term) -> BNF_Def.bnf list -> thm ->
- thm list -> BNF_Def.bnf list -> typ list -> typ list -> typ list list list -> int list ->
+ val derive_coinduct_thms_for_types: Proof.context -> bool -> (term -> term) -> BNF_Def.bnf list ->
+ thm -> thm list -> BNF_Def.bnf list -> typ list -> typ list -> typ list list list -> int list ->
thm list -> thm list -> (thm -> thm) -> thm list list -> Ctr_Sugar.ctr_sugar list ->
- Proof.context -> (thm list * thm) list
- val derive_coinduct_corecs_thms_for_types: BNF_Def.bnf list ->
+ (thm list * thm) list
+ val derive_coinduct_corecs_thms_for_types: Proof.context -> BNF_Def.bnf list ->
string * term list * term list list
* (((term list list * term list list * term list list list list * term list list list list)
* term list list list) * typ list) ->
thm -> thm list -> thm list -> thm list -> BNF_Def.bnf list -> typ list -> typ list ->
typ list -> typ list list list -> int list list -> int list list -> int list -> thm list ->
thm list -> (thm -> thm) -> thm list list -> Ctr_Sugar.ctr_sugar list -> term list ->
- thm list -> Proof.context -> gfp_sugar_thms
+ thm list -> gfp_sugar_thms
val co_datatypes: BNF_Util.fp_kind -> (mixfix list -> binding list -> binding list ->
binding list -> binding list list -> binding list -> (string * sort) list ->
@@ -524,7 +533,7 @@
Option.map (rapp u o rapp t) (choose_binary_fun fs (fastype_of t, fastype_of u));
fun build_the_rel ctxt Rs Ts A B =
- build_rel [] ctxt Ts (the o choose_binary_fun Rs) (A, B);
+ build_rel [] ctxt Ts [] (the o choose_binary_fun Rs) (A, B);
fun build_rel_app ctxt Rs Ts t u =
build_the_rel ctxt Rs Ts (fastype_of t) (fastype_of u) $ t $ u;
@@ -568,6 +577,19 @@
val mk_ctor = mk_ctor_or_dtor range_type;
val mk_dtor = mk_ctor_or_dtor domain_type;
+fun mk_bnf_sets bnf =
+ let
+ val Type (T_name, Us) = T_of_bnf bnf;
+ val lives = lives_of_bnf bnf;
+ val sets = sets_of_bnf bnf;
+ fun mk_set U =
+ (case find_index (curry (op =) U) lives of
+ ~1 => Term.dummy
+ | i => nth sets i);
+ in
+ (T_name, map mk_set Us)
+ end;
+
fun mk_xtor_co_recs thy fp fpTs Cs ts0 =
let
val nn = length fpTs;
@@ -614,10 +636,10 @@
fun define_ctrs_dtrs_for_type fp_b_name fpT ctor dtor ctor_dtor dtor_ctor n ks abs ctr_bindings
ctr_mixfixes ctr_Tss lthy =
let
- val ctr_absT = domain_type (fastype_of ctor);
+ val ctor_absT = domain_type (fastype_of ctor);
val (((w, xss), u'), _) = lthy
- |> yield_singleton (mk_Frees "w") ctr_absT
+ |> yield_singleton (mk_Frees "w") ctor_absT
||>> mk_Freess "x" ctr_Tss
||>> yield_singleton Variable.variant_fixes fp_b_name;
@@ -631,13 +653,13 @@
val vars = Variable.add_free_names lthy goal [];
in
Goal.prove_sorry lthy vars [] goal (fn {context = ctxt, ...} =>
- mk_ctor_iff_dtor_tac ctxt (map (SOME o Thm.ctyp_of lthy) [ctr_absT, fpT])
+ mk_ctor_iff_dtor_tac ctxt (map (SOME o Thm.ctyp_of lthy) [ctor_absT, fpT])
(Thm.cterm_of lthy ctor) (Thm.cterm_of lthy dtor) ctor_dtor dtor_ctor)
|> Thm.close_derivation
end;
val ctr_rhss =
- map2 (fn k => fn xs => fold_rev Term.lambda xs (ctor $ mk_absumprod ctr_absT abs n k xs))
+ map2 (fn k => fn xs => fold_rev Term.lambda xs (ctor $ mk_absumprod ctor_absT abs n k xs))
ks xss;
val ((raw_ctrs, raw_ctr_defs), (lthy, lthy_old)) = lthy
@@ -841,7 +863,7 @@
fun mk_arg (x as Free (_, T)) (Free (_, U)) =
if T = U then x
- else build_map lthy [] (the o AList.lookup (op =) ABfs) (T, U) $ x;
+ else build_map lthy [] [] (the o AList.lookup (op =) ABfs) (T, U) $ x;
val xs' = map2 mk_arg xs ys;
in
@@ -1223,7 +1245,7 @@
val lhsT = fastype_of lhs;
val map_rhsT =
map_atyps (perhaps (AList.lookup (op =) (map swap live_AsBs))) lhsT;
- val map_rhs = build_map lthy []
+ val map_rhs = build_map lthy [] []
(the o (AList.lookup (op =) (live_AsBs ~~ fs))) (map_rhsT, lhsT);
val rhs = (case map_rhs of
Const (@{const_name id}, _) => selA $ ta
@@ -1520,7 +1542,7 @@
val cpss = map2 (map o rapp) cs pss;
- fun build_sum_inj mk_inj = build_map ctxt [] (uncurry mk_inj o dest_sumT o snd);
+ fun build_sum_inj mk_inj = build_map ctxt [] [] (uncurry mk_inj o dest_sumT o snd);
fun build_dtor_corec_arg _ [] [cg] = cg
| build_dtor_corec_arg T [cq] [cg, cg'] =
@@ -1604,6 +1626,80 @@
@{map 5} mk_preds_getterss_join cs cpss f_absTs abss cqgsss)))
end;
+fun mk_induct_raw_prem_prems names_ctxt Xss setss_fp_nesting (x as Free (s, Type (T_name, Ts0)))
+ (Type (_, Xs_Ts0)) =
+ (case AList.lookup (op =) setss_fp_nesting T_name of
+ NONE => []
+ | SOME raw_sets0 =>
+ let
+ val (Xs_Ts, (Ts, raw_sets)) =
+ filter (exists_subtype_in (flat Xss) o fst) (Xs_Ts0 ~~ (Ts0 ~~ raw_sets0))
+ |> split_list ||> split_list;
+ val sets = map (mk_set Ts0) raw_sets;
+ val (ys, names_ctxt') = names_ctxt |> mk_Frees s Ts;
+ val xysets = map (pair x) (ys ~~ sets);
+ val ppremss = map2 (mk_induct_raw_prem_prems names_ctxt' Xss setss_fp_nesting) ys Xs_Ts;
+ in
+ flat (map2 (map o apfst o cons) xysets ppremss)
+ end)
+ | mk_induct_raw_prem_prems _ Xss _ (x as Free (_, Type _)) X =
+ [([], (find_index (fn Xs => member (op =) Xs X) Xss + 1, x))]
+ | mk_induct_raw_prem_prems _ _ _ _ _ = [];
+
+fun mk_induct_raw_prem alter_x names_ctxt Xss setss_fp_nesting p ctr ctr_Ts ctrXs_Ts =
+ let
+ val (xs, names_ctxt') = names_ctxt |> mk_Frees "x" ctr_Ts;
+ val pprems =
+ flat (map2 (mk_induct_raw_prem_prems names_ctxt' Xss setss_fp_nesting) xs ctrXs_Ts);
+ val y = Term.list_comb (ctr, map alter_x xs);
+ val p' = enforce_type names_ctxt domain_type (fastype_of y) p;
+ in (xs, pprems, HOLogic.mk_Trueprop (p' $ y)) end;
+
+fun close_induct_prem_prem nn ps xs t =
+ fold_rev Logic.all (map Free (drop (nn + length xs)
+ (rev (Term.add_frees t (map dest_Free xs @ map_filter (try dest_Free) ps))))) t;
+
+fun finish_induct_prem_prem ctxt nn ps xs (xysets, (j, x)) =
+ let val p' = enforce_type ctxt domain_type (fastype_of x) (nth ps (j - 1)) in
+ close_induct_prem_prem nn ps xs (Logic.list_implies (map (fn (x', (y, set)) =>
+ mk_Trueprop_mem (y, set $ x')) xysets,
+ HOLogic.mk_Trueprop (p' $ x)))
+ end;
+
+fun finish_induct_prem ctxt nn ps (xs, raw_pprems, concl) =
+ fold_rev Logic.all xs (Logic.list_implies
+ (map (finish_induct_prem_prem ctxt nn ps xs) raw_pprems, concl));
+
+fun mk_coinduct_prem_ctr_concls ctxt Xss fpTss rs' n k udisc usels vdisc vsels ctrXs_Ts =
+ let
+ fun build_the_rel T Xs_T =
+ build_rel [] ctxt [] [] (fn (T, X) =>
+ nth rs' (find_index (fn Xs => member (op =) Xs X) Xss)
+ |> enforce_type ctxt domain_type T)
+ (T, Xs_T)
+ |> Term.subst_atomic_types (flat Xss ~~ flat fpTss);
+ fun build_rel_app usel vsel Xs_T =
+ fold rapp [usel, vsel] (build_the_rel (fastype_of usel) Xs_T);
+ in
+ (if k = n then [] else [HOLogic.mk_eq (udisc, vdisc)]) @
+ (if null usels then
+ []
+ else
+ [Library.foldr HOLogic.mk_imp (if n = 1 then [] else [udisc, vdisc],
+ Library.foldr1 HOLogic.mk_conj (@{map 3} build_rel_app usels vsels ctrXs_Ts))])
+ end;
+
+fun mk_coinduct_prem_concl ctxt Xss fpTss rs' n udiscs uselss vdiscs vselss ctrXs_Tss =
+ @{map 6} (mk_coinduct_prem_ctr_concls ctxt Xss fpTss rs' n)
+ (1 upto n) udiscs uselss vdiscs vselss ctrXs_Tss
+ |> flat |> Library.foldr1 HOLogic.mk_conj
+ handle List.Empty => @{term True};
+
+fun mk_coinduct_prem ctxt Xss fpTss rs' uvr u v n udiscs uselss vdiscs vselss ctrXs_Tss =
+ fold_rev Logic.all [u, v] (Logic.mk_implies (HOLogic.mk_Trueprop uvr,
+ HOLogic.mk_Trueprop (mk_coinduct_prem_concl ctxt Xss fpTss rs' n udiscs uselss vdiscs vselss
+ ctrXs_Tss)));
+
fun postproc_co_induct ctxt nn prop prop_conj =
Drule.zero_var_indexes
#> `(conj_dests nn)
@@ -1616,7 +1712,7 @@
let val induct_cases = quasi_unambiguous_case_names (maps (map name_of_ctr) ctrss);
in [Attrib.case_names induct_cases] end;
-fun derive_rel_induct_thms_for_types lthy nn fpA_Ts As Bs ctrAss ctrAs_Tsss exhausts ctor_rel_induct
+fun derive_rel_induct_thms_for_types ctxt nn fpA_Ts As Bs ctrAss ctrAs_Tsss exhausts ctor_rel_induct
ctor_defss ctor_injects pre_rel_defs abs_inverses live_nesting_rel_eqs =
let
val B_ify_T = Term.typ_subst_atomic (As ~~ Bs);
@@ -1626,7 +1722,7 @@
val ctrBs_Tsss = map (map (map B_ify_T)) ctrAs_Tsss;
val ctrBss = map (map B_ify) ctrAss;
- val ((((Rs, IRs), ctrAsss), ctrBsss), names_lthy) = lthy
+ val ((((Rs, IRs), ctrAsss), ctrBsss), names_ctxt) = ctxt
|> mk_Frees "R" (map2 mk_pred2T As Bs)
||>> mk_Frees "IR" (map2 mk_pred2T fpA_Ts fpB_Ts)
||>> mk_Freesss "a" ctrAs_Tsss
@@ -1636,30 +1732,30 @@
let
fun mk_prem ctrA ctrB argAs argBs =
fold_rev Logic.all (argAs @ argBs) (fold_rev (curry Logic.mk_implies)
- (map2 (HOLogic.mk_Trueprop oo build_rel_app names_lthy (Rs @ IRs) fpA_Ts) argAs argBs)
- (HOLogic.mk_Trueprop (build_rel_app names_lthy (Rs @ IRs) fpA_Ts
+ (map2 (HOLogic.mk_Trueprop oo build_rel_app names_ctxt (Rs @ IRs) fpA_Ts) argAs argBs)
+ (HOLogic.mk_Trueprop (build_rel_app names_ctxt (Rs @ IRs) fpA_Ts
(Term.list_comb (ctrA, argAs)) (Term.list_comb (ctrB, argBs)))));
in
flat (@{map 4} (@{map 4} mk_prem) ctrAss ctrBss ctrAsss ctrBsss)
end;
val goal = HOLogic.mk_Trueprop (Library.foldr1 HOLogic.mk_conj (map2 mk_leq
- (map2 (build_the_rel lthy (Rs @ IRs) []) fpA_Ts fpB_Ts) IRs));
- val vars = Variable.add_free_names lthy goal [];
+ (map2 (build_the_rel ctxt (Rs @ IRs) []) fpA_Ts fpB_Ts) IRs));
+ val vars = Variable.add_free_names ctxt goal [];
val rel_induct0_thm =
- Goal.prove_sorry lthy vars prems goal (fn {context = ctxt, prems} =>
+ Goal.prove_sorry ctxt vars prems goal (fn {context = ctxt, prems} =>
mk_rel_induct0_tac ctxt ctor_rel_induct prems (map (Thm.cterm_of ctxt) IRs) exhausts
ctor_defss ctor_injects pre_rel_defs abs_inverses live_nesting_rel_eqs)
|> Thm.close_derivation;
in
- (postproc_co_induct lthy nn @{thm predicate2D} @{thm predicate2D_conj} rel_induct0_thm,
+ (postproc_co_induct ctxt nn @{thm predicate2D} @{thm predicate2D_conj} rel_induct0_thm,
mk_induct_attrs ctrAss)
end;
fun derive_induct_recs_thms_for_types plugins pre_bnfs rec_args_typess ctor_induct ctor_rec_thms
- live_nesting_bnfs fp_nesting_bnfs fpTs Cs Xs ctrXs_Tsss fp_abs_inverses fp_type_definitions
- abs_inverses ctrss ctr_defss recs rec_defs lthy =
+ live_nesting_bnfs fp_nesting_bnfs fpTs Cs Xs ctrXs_Tsss pre_abs_inverses pre_type_definitions
+ abs_inverses ctrss ctr_defss recs rec_defs ctxt =
let
val ctr_Tsss = map (map (binder_types o fastype_of)) ctrss;
@@ -1675,81 +1771,33 @@
val fp_b_names = map base_name_of_typ fpTs;
- val ((((ps, ps'), xsss), us'), names_lthy) = lthy
- |> mk_Frees' "P" (map mk_pred1T fpTs)
+ val (((ps, xsss), us'), names_ctxt) = ctxt
+ |> mk_Frees "P" (map mk_pred1T fpTs)
||>> mk_Freesss "x" ctr_Tsss
||>> Variable.variant_fixes fp_b_names;
val us = map2 (curry Free) us' fpTs;
- fun mk_sets bnf =
- let
- val Type (T_name, Us) = T_of_bnf bnf;
- val lives = lives_of_bnf bnf;
- val sets = sets_of_bnf bnf;
- fun mk_set U =
- (case find_index (curry (op =) U) lives of
- ~1 => Term.dummy
- | i => nth sets i);
- in
- (T_name, map mk_set Us)
- end;
-
- val setss_fp_nesting = map mk_sets fp_nesting_bnfs;
+ val setss_fp_nesting = map mk_bnf_sets fp_nesting_bnfs;
val (induct_thms, induct_thm) =
let
- fun mk_raw_prem_prems _ (x as Free (_, Type _)) (X as TFree _) =
- [([], (find_index (curry (op =) X) Xs + 1, x))]
- | mk_raw_prem_prems names_lthy (x as Free (s, Type (T_name, Ts0))) (Type (_, Xs_Ts0)) =
- (case AList.lookup (op =) setss_fp_nesting T_name of
- NONE => []
- | SOME raw_sets0 =>
- let
- val (Xs_Ts, (Ts, raw_sets)) =
- filter (exists_subtype_in Xs o fst) (Xs_Ts0 ~~ (Ts0 ~~ raw_sets0))
- |> split_list ||> split_list;
- val sets = map (mk_set Ts0) raw_sets;
- val (ys, names_lthy') = names_lthy |> mk_Frees s Ts;
- val xysets = map (pair x) (ys ~~ sets);
- val ppremss = map2 (mk_raw_prem_prems names_lthy') ys Xs_Ts;
- in
- flat (map2 (map o apfst o cons) xysets ppremss)
- end)
- | mk_raw_prem_prems _ _ _ = [];
-
- fun close_prem_prem xs t =
- fold_rev Logic.all (map Free (drop (nn + length xs)
- (rev (Term.add_frees t (map dest_Free xs @ ps'))))) t;
-
- fun mk_prem_prem xs (xysets, (j, x)) =
- close_prem_prem xs (Logic.list_implies (map (fn (x', (y, set)) =>
- mk_Trueprop_mem (y, set $ x')) xysets,
- HOLogic.mk_Trueprop (nth ps (j - 1) $ x)));
-
- fun mk_raw_prem phi ctr ctr_Ts ctrXs_Ts =
- let
- val (xs, names_lthy') = names_lthy |> mk_Frees "x" ctr_Ts;
- val pprems = flat (map2 (mk_raw_prem_prems names_lthy') xs ctrXs_Ts);
- in (xs, pprems, HOLogic.mk_Trueprop (phi $ Term.list_comb (ctr, xs))) end;
-
- fun mk_prem (xs, raw_pprems, concl) =
- fold_rev Logic.all xs (Logic.list_implies (map (mk_prem_prem xs) raw_pprems, concl));
-
- val raw_premss = @{map 4} (@{map 3} o mk_raw_prem) ps ctrss ctr_Tsss ctrXs_Tsss;
-
+ val raw_premss = @{map 4} (@{map 3}
+ o mk_induct_raw_prem I names_ctxt (map single Xs) setss_fp_nesting)
+ ps ctrss ctr_Tsss ctrXs_Tsss;
+ val concl =
+ HOLogic.mk_Trueprop (Library.foldr1 HOLogic.mk_conj (map2 (curry (op $)) ps us));
val goal =
- Library.foldr (Logic.list_implies o apfst (map mk_prem)) (raw_premss,
- HOLogic.mk_Trueprop (Library.foldr1 HOLogic.mk_conj (map2 (curry (op $)) ps us)));
- val vars = Variable.add_free_names lthy goal [];
-
+ Library.foldr (Logic.list_implies o apfst (map (finish_induct_prem ctxt nn ps)))
+ (raw_premss, concl);
+ val vars = Variable.add_free_names ctxt goal [];
val kksss = map (map (map (fst o snd) o #2)) raw_premss;
- val ctor_induct' = ctor_induct OF (map2 mk_absumprodE fp_type_definitions mss);
+ val ctor_induct' = ctor_induct OF (map2 mk_absumprodE pre_type_definitions mss);
val thm =
- Goal.prove_sorry lthy vars [] goal (fn {context = ctxt, ...} =>
- mk_induct_tac ctxt nn ns mss kksss (flat ctr_defss) ctor_induct' fp_abs_inverses
+ Goal.prove_sorry ctxt vars [] goal (fn {context = ctxt, ...} =>
+ mk_induct_tac ctxt nn ns mss kksss (flat ctr_defss) ctor_induct' pre_abs_inverses
abs_inverses fp_nesting_set_maps pre_set_defss)
|> Thm.close_derivation;
in
@@ -1781,7 +1829,7 @@
indexify (perhaps (try (snd o HOLogic.dest_prodT)) o snd) Cs
(fn kk => fn TU => maybe_tick TU (nth us kk) (nth frecs kk));
in
- build_map lthy [] build_simple (T, U) $ x
+ build_map ctxt [] [] build_simple (T, U) $ x
end;
val fxsss = map2 (map2 (flat_rec_arg_args oo map2 (map o build_rec))) xsss x_Tssss;
@@ -1789,10 +1837,10 @@
val tacss = @{map 4} (map ooo
mk_rec_tac pre_map_defs (fp_nesting_map_ident0s @ live_nesting_map_ident0s) rec_defs)
- ctor_rec_thms fp_abs_inverses abs_inverses ctr_defss;
+ ctor_rec_thms pre_abs_inverses abs_inverses ctr_defss;
fun prove goal tac =
- Goal.prove_sorry lthy [] [] goal (tac o #context)
+ Goal.prove_sorry ctxt [] [] goal (tac o #context)
|> Thm.close_derivation;
in
map2 (map2 prove) goalss tacss
@@ -1837,7 +1885,7 @@
(coinduct_attrs, common_coinduct_attrs)
end;
-fun derive_rel_coinduct_thms_for_types lthy nn fpA_Ts ns As Bs mss (ctr_sugars : ctr_sugar list)
+fun derive_rel_coinduct_thms_for_types ctxt nn fpA_Ts ns As Bs mss (ctr_sugars : ctr_sugar list)
abs_inverses abs_injects ctor_injects dtor_ctors rel_pre_defs ctor_defss dtor_rel_coinduct
live_nesting_rel_eqs =
let
@@ -1847,14 +1895,14 @@
val (Rs, IRs, fpAs, fpBs, _) =
let
val fp_names = map base_name_of_typ fpA_Ts;
- val ((((Rs, IRs), fpAs_names), fpBs_names), names_lthy) = lthy
+ val ((((Rs, IRs), fpAs_names), fpBs_names), names_ctxt) = ctxt
|> mk_Frees "R" (map2 mk_pred2T As Bs)
||>> mk_Frees "IR" (map2 mk_pred2T fpA_Ts fpB_Ts)
||>> Variable.variant_fixes fp_names
||>> Variable.variant_fixes (map (suffix "'") fp_names);
in
(Rs, IRs, map2 (curry Free) fpAs_names fpA_Ts, map2 (curry Free) fpBs_names fpB_Ts,
- names_lthy)
+ names_ctxt)
end;
val ((discA_tss, selA_tsss), (discB_tss, selB_tsss)) =
@@ -1880,7 +1928,7 @@
[Library.foldr HOLogic.mk_imp
(if n = 1 then [] else [discA_t, discB_t],
Library.foldr1 HOLogic.mk_conj
- (map2 (build_rel_app lthy (Rs @ IRs) fpA_Ts) selA_ts selB_ts))]);
+ (map2 (build_rel_app ctxt (Rs @ IRs) fpA_Ts) selA_ts selB_ts))]);
fun mk_prem_concl n discA_ts selA_tss discB_ts selB_tss =
Library.foldr1 HOLogic.mk_conj (flat (@{map 5} (mk_prem_ctr_concls n)
@@ -1895,22 +1943,22 @@
end;
val goal = HOLogic.mk_Trueprop (Library.foldr1 HOLogic.mk_conj (map2 mk_leq
- IRs (map2 (build_the_rel lthy (Rs @ IRs) []) fpA_Ts fpB_Ts)));
- val vars = Variable.add_free_names lthy goal [];
+ IRs (map2 (build_the_rel ctxt (Rs @ IRs) []) fpA_Ts fpB_Ts)));
+ val vars = Variable.add_free_names ctxt goal [];
val rel_coinduct0_thm =
- Goal.prove_sorry lthy vars prems goal (fn {context = ctxt, prems} =>
+ Goal.prove_sorry ctxt vars prems goal (fn {context = ctxt, prems} =>
mk_rel_coinduct0_tac ctxt dtor_rel_coinduct (map (Thm.cterm_of ctxt) IRs) prems
(map #exhaust ctr_sugars) (map (flat o #disc_thmss) ctr_sugars)
(map (flat o #sel_thmss) ctr_sugars) ctor_defss dtor_ctors ctor_injects abs_injects
rel_pre_defs abs_inverses live_nesting_rel_eqs)
|> Thm.close_derivation;
in
- (postproc_co_induct lthy nn @{thm predicate2D} @{thm predicate2D_conj} rel_coinduct0_thm,
+ (postproc_co_induct ctxt nn @{thm predicate2D} @{thm predicate2D_conj} rel_coinduct0_thm,
mk_coinduct_attrs fpA_Ts (map #ctrs ctr_sugars) (map #discs ctr_sugars) mss)
end;
-fun derive_set_induct_thms_for_types lthy nn fpTs ctrss setss dtor_set_inducts exhausts set_pre_defs
+fun derive_set_induct_thms_for_types ctxt nn fpTs ctrss setss dtor_set_inducts exhausts set_pre_defs
ctor_defs dtor_ctors Abs_pre_inverses =
let
fun mk_prems A Ps ctr_args t ctxt =
@@ -1979,9 +2027,9 @@
|>> apfst (apsnd (Library.foldr1 HOLogic.mk_conj))
|>> apsnd flat;
- val vars = fold (Variable.add_free_names lthy) (concl :: prems) [];
+ val vars = fold (Variable.add_free_names ctxt) (concl :: prems) [];
val thm =
- Goal.prove_sorry lthy vars prems (HOLogic.mk_Trueprop concl)
+ Goal.prove_sorry ctxt vars prems (HOLogic.mk_Trueprop concl)
(fn {context = ctxt, prems} =>
mk_set_induct0_tac ctxt (map (Thm.cterm_of ctxt'') Ps) prems dtor_set_inducts
exhausts set_pre_defs ctor_defs dtor_ctors Abs_pre_inverses)
@@ -1994,7 +2042,7 @@
end
val consumes_attr = Attrib.consumes 1;
in
- map (mk_thm lthy fpTs ctrss
+ map (mk_thm ctxt fpTs ctrss
#> nn = 1 ? map_prod (fn thm => rotate_prems ~1 (thm RS bspec)) (cons consumes_attr))
(transpose setss)
end;
@@ -2018,9 +2066,9 @@
|> unfold_thms ctxt unfolds
end;
-fun derive_coinduct_thms_for_types strong alter_r pre_bnfs dtor_coinduct dtor_ctors
- live_nesting_bnfs fpTs Xs ctrXs_Tsss ns fp_abs_inverses abs_inverses mk_vimage2p ctr_defss
- (ctr_sugars : ctr_sugar list) ctxt =
+fun derive_coinduct_thms_for_types ctxt strong alter_r pre_bnfs dtor_coinduct dtor_ctors
+ live_nesting_bnfs fpTs Xs ctrXs_Tsss ns pre_abs_inverses abs_inverses mk_vimage2p ctr_defss
+ (ctr_sugars : ctr_sugar list) =
let
val nn = length pre_bnfs;
@@ -2054,45 +2102,23 @@
@{map 4} (fn u => fn v => fn uvr => fn uv_eq =>
fold_rev Term.lambda [u, v] (HOLogic.mk_disj (uvr, uv_eq))) us vs uvrs uv_eqs;
- fun build_the_rel rs' T Xs_T =
- build_rel [] ctxt [] (fn (_, X) => nth rs' (find_index (curry (op =) X) Xs)) (T, Xs_T)
- |> Term.subst_atomic_types (Xs ~~ fpTs);
-
- fun build_rel_app rs' usel vsel Xs_T =
- fold rapp [usel, vsel] (build_the_rel rs' (fastype_of usel) Xs_T);
-
- fun mk_prem_ctr_concls rs' n k udisc usels vdisc vsels ctrXs_Ts =
- (if k = n then [] else [HOLogic.mk_eq (udisc, vdisc)]) @
- (if null usels then
- []
- else
- [Library.foldr HOLogic.mk_imp (if n = 1 then [] else [udisc, vdisc],
- Library.foldr1 HOLogic.mk_conj (@{map 3} (build_rel_app rs') usels vsels ctrXs_Ts))]);
-
- fun mk_prem_concl rs' n udiscs uselss vdiscs vselss ctrXs_Tss =
- flat (@{map 6} (mk_prem_ctr_concls rs' n) (1 upto n) udiscs uselss vdiscs vselss ctrXs_Tss)
- |> Library.foldr1 HOLogic.mk_conj
- handle List.Empty => @{term True};
-
- fun mk_prem rs' uvr u v n udiscs uselss vdiscs vselss ctrXs_Tss =
- fold_rev Logic.all [u, v] (Logic.mk_implies (HOLogic.mk_Trueprop uvr,
- HOLogic.mk_Trueprop (mk_prem_concl rs' n udiscs uselss vdiscs vselss ctrXs_Tss)));
-
val concl =
HOLogic.mk_Trueprop (Library.foldr1 HOLogic.mk_conj
(@{map 3} (fn uvr => fn u => fn v => HOLogic.mk_imp (uvr, HOLogic.mk_eq (u, v)))
- uvrs us vs));
+ uvrs us vs))
fun mk_goal rs0' =
- Logic.list_implies (@{map 9} (mk_prem (map alter_r rs0')) uvrs us vs ns udiscss uselsss
- vdiscss vselsss ctrXs_Tsss, concl);
+ Logic.list_implies (@{map 9} (mk_coinduct_prem ctxt (map single Xs) (map single fpTs)
+ (map alter_r rs0'))
+ uvrs us vs ns udiscss uselsss vdiscss vselsss ctrXs_Tsss,
+ concl);
val goals = map mk_goal ([rs] @ (if strong then [strong_rs] else []));
fun prove dtor_coinduct' goal =
Variable.add_free_names ctxt goal []
|> (fn vars => Goal.prove_sorry ctxt vars [] goal (fn {context = ctxt, ...} =>
- mk_coinduct_tac ctxt live_nesting_rel_eqs nn ns dtor_coinduct' pre_rel_defs fp_abs_inverses
+ mk_coinduct_tac ctxt live_nesting_rel_eqs nn ns dtor_coinduct' pre_rel_defs pre_abs_inverses
abs_inverses dtor_ctors exhausts ctr_defss disc_thmsss sel_thmsss))
|> Thm.close_derivation;
@@ -2107,10 +2133,10 @@
dtor_coinducts goals
end;
-fun derive_coinduct_corecs_thms_for_types pre_bnfs (x, cs, cpss, (((pgss, _, _, _), cqgsss), _))
- dtor_coinduct dtor_injects dtor_ctors dtor_corec_thms live_nesting_bnfs fpTs Cs Xs ctrXs_Tsss
- kss mss ns fp_abs_inverses abs_inverses mk_vimage2p ctr_defss (ctr_sugars : ctr_sugar list)
- corecs corec_defs ctxt =
+fun derive_coinduct_corecs_thms_for_types ctxt pre_bnfs
+ (x, cs, cpss, (((pgss, _, _, _), cqgsss), _)) dtor_coinduct dtor_injects dtor_ctors
+ dtor_corec_thms live_nesting_bnfs fpTs Cs Xs ctrXs_Tsss kss mss ns pre_abs_inverses abs_inverses
+ mk_vimage2p ctr_defss (ctr_sugars : ctr_sugar list) corecs corec_defs =
let
fun mk_ctor_dtor_corec_thm dtor_inject dtor_ctor corec =
iffD1 OF [dtor_inject, trans OF [corec, dtor_ctor RS sym]];
@@ -2130,9 +2156,9 @@
val discIss = map #discIs ctr_sugars;
val sel_thmsss = map #sel_thmss ctr_sugars;
- val coinduct_thms_pairs = derive_coinduct_thms_for_types true I pre_bnfs dtor_coinduct
- dtor_ctors live_nesting_bnfs fpTs Xs ctrXs_Tsss ns fp_abs_inverses abs_inverses mk_vimage2p
- ctr_defss ctr_sugars ctxt;
+ val coinduct_thms_pairs = derive_coinduct_thms_for_types ctxt true I pre_bnfs dtor_coinduct
+ dtor_ctors live_nesting_bnfs fpTs Xs ctrXs_Tsss ns pre_abs_inverses abs_inverses mk_vimage2p
+ ctr_defss ctr_sugars;
fun mk_maybe_not pos = not pos ? HOLogic.mk_not;
@@ -2166,7 +2192,7 @@
indexify fst (map2 (curry mk_sumT) fpTs Cs)
(fn kk => fn _ => tack (nth cs kk, nth us kk) (nth gcorecs kk));
in
- build_map ctxt [] build_simple (T, U) $ cqg
+ build_map ctxt [] [] build_simple (T, U) $ cqg
end
else
cqg
@@ -2579,7 +2605,7 @@
fun mk_rec_arg_arg (x as Free (_, T)) =
let val U = B_ify_T T in
- if T = U then x else build_map lthy [] (the o AList.lookup (op =) ABfs) (T, U) $ x
+ if T = U then x else build_map lthy [] [] (the o AList.lookup (op =) ABfs) (T, U) $ x
end;
fun mk_rec_o_map_arg rec_arg_T h =
@@ -2732,8 +2758,10 @@
val T = range_type (fastype_of g);
val U = range_type corec_argB_T;
in
- if T = U then g
- else HOLogic.mk_comp (build_map lthy2 [] (the o AList.lookup (op =) ABfs) (T, U), g)
+ if T = U then
+ g
+ else
+ HOLogic.mk_comp (build_map lthy2 [] [] (the o AList.lookup (op =) ABfs) (T, U), g)
end;
fun mk_map_o_corec_rhs corecx =
@@ -2769,9 +2797,9 @@
(coinduct_attrs, common_coinduct_attrs)),
corec_thmss, corec_disc_thmss,
(corec_disc_iff_thmss, corec_disc_iff_attrs), (corec_sel_thmsss, corec_sel_attrs)) =
- derive_coinduct_corecs_thms_for_types pre_bnfs (the corecs_args_types) xtor_co_induct
+ derive_coinduct_corecs_thms_for_types lthy pre_bnfs (the corecs_args_types) xtor_co_induct
dtor_injects dtor_ctors xtor_co_rec_thms live_nesting_bnfs fpTs Cs Xs ctrXs_Tsss kss mss
- ns abs_inverses abs_inverses I ctr_defss ctr_sugars corecs corec_defs lthy;
+ ns abs_inverses abs_inverses I ctr_defss ctr_sugars corecs corec_defs;
fun distinct_prems ctxt thm =
Goal.prove (*no sorry*) ctxt [] []
--- a/src/HOL/Tools/BNF/bnf_fp_def_sugar_tactics.ML Wed Jan 04 14:26:08 2017 +0100
+++ b/src/HOL/Tools/BNF/bnf_fp_def_sugar_tactics.ML Wed Jan 04 14:26:19 2017 +0100
@@ -10,7 +10,10 @@
sig
val sumprod_thms_rel: thm list
+ val co_induct_inst_as_projs_tac: Proof.context -> int -> tactic
val mk_case_transfer_tac: Proof.context -> thm -> thm list -> tactic
+ val mk_coinduct_discharge_prem_tac: Proof.context -> thm list -> thm list -> int -> int -> int ->
+ thm -> thm -> thm -> thm -> thm -> thm list -> thm list list -> thm list list -> int -> tactic
val mk_coinduct_tac: Proof.context -> thm list -> int -> int list -> thm -> thm list ->
thm list -> thm list -> thm list -> thm list -> thm list list -> thm list list list ->
thm list list list -> tactic
@@ -27,6 +30,8 @@
val mk_disc_transfer_tac: Proof.context -> thm -> thm -> thm list -> tactic
val mk_exhaust_tac: Proof.context -> int -> thm list -> thm -> thm -> tactic
val mk_half_distinct_tac: Proof.context -> thm -> thm -> thm list -> tactic
+ val mk_induct_discharge_prem_tac: Proof.context -> int -> int -> thm list -> thm list ->
+ thm list -> thm list -> int -> int -> int list -> tactic
val mk_induct_tac: Proof.context -> int -> int list -> int list list -> int list list list ->
thm list -> thm -> thm list -> thm list -> thm list -> thm list list -> tactic
val mk_inject_tac: Proof.context -> thm -> thm -> thm -> tactic
@@ -176,10 +181,10 @@
rec_o_map_simps) ctxt))
end;
-fun mk_rec_tac pre_map_defs map_ident0s rec_defs ctor_rec fp_abs_inverse abs_inverse ctr_def ctxt =
+fun mk_rec_tac pre_map_defs map_ident0s rec_defs ctor_rec pre_abs_inverse abs_inverse ctr_def ctxt =
HEADGOAL ((if is_def_looping ctr_def then subst_tac ctxt NONE
else SELECT_GOAL o unfold_thms_tac ctxt) [ctr_def]) THEN
- unfold_thms_tac ctxt (ctor_rec :: fp_abs_inverse :: abs_inverse :: rec_defs @
+ unfold_thms_tac ctxt (ctor_rec :: pre_abs_inverse :: abs_inverse :: rec_defs @
pre_map_defs @ map_ident0s @ rec_unfold_thms) THEN HEADGOAL (rtac ctxt refl);
fun mk_rec_transfer_tac ctxt nn ns actives passives xssss rec_defs ctor_rec_transfers rel_pre_T_defs
@@ -313,15 +318,15 @@
resolve_tac ctxt @{thms disjI1 disjI2}) THEN'
(rtac ctxt refl ORELSE' assume_tac ctxt ORELSE' rtac ctxt @{thm singletonI});
-fun mk_induct_leverage_prem_prems_tac ctxt nn kks fp_abs_inverses abs_inverses set_maps
+fun mk_induct_leverage_prem_prems_tac ctxt nn kks pre_abs_inverses abs_inverses set_maps
pre_set_defs =
HEADGOAL (EVERY' (maps (fn kk => [select_prem_tac ctxt nn (dtac ctxt meta_spec) kk,
etac ctxt meta_mp,
- SELECT_GOAL (unfold_thms_tac ctxt (pre_set_defs @ fp_abs_inverses @ abs_inverses @ set_maps @
+ SELECT_GOAL (unfold_thms_tac ctxt (pre_set_defs @ pre_abs_inverses @ abs_inverses @ set_maps @
sumprod_thms_set)),
solve_prem_prem_tac ctxt]) (rev kks)));
-fun mk_induct_discharge_prem_tac ctxt nn n fp_abs_inverses abs_inverses set_maps pre_set_defs m k
+fun mk_induct_discharge_prem_tac ctxt nn n pre_abs_inverses abs_inverses set_maps pre_set_defs m k
kks =
let val r = length kks in
HEADGOAL (EVERY' [select_prem_tac ctxt n (rotate_tac 1) k, rotate_tac ~1, hyp_subst_tac ctxt,
@@ -329,11 +334,11 @@
EVERY [REPEAT_DETERM_N r
(HEADGOAL (rotate_tac ~1 THEN' dtac ctxt meta_mp THEN' rotate_tac 1) THEN prefer_tac 2),
if r > 0 then ALLGOALS (Goal.norm_hhf_tac ctxt) else all_tac, HEADGOAL (assume_tac ctxt),
- mk_induct_leverage_prem_prems_tac ctxt nn kks fp_abs_inverses abs_inverses set_maps
+ mk_induct_leverage_prem_prems_tac ctxt nn kks pre_abs_inverses abs_inverses set_maps
pre_set_defs]
end;
-fun mk_induct_tac ctxt nn ns mss kkss ctr_defs ctor_induct' fp_abs_inverses abs_inverses set_maps
+fun mk_induct_tac ctxt nn ns mss kksss ctr_defs ctor_induct' pre_abs_inverses abs_inverses set_maps
pre_set_defss =
let val n = Integer.sum ns in
(if exists is_def_looping ctr_defs then
@@ -342,18 +347,18 @@
unfold_thms_tac ctxt ctr_defs) THEN
HEADGOAL (rtac ctxt ctor_induct') THEN co_induct_inst_as_projs_tac ctxt 0 THEN
EVERY (@{map 4} (EVERY oooo @{map 3} o
- mk_induct_discharge_prem_tac ctxt nn n fp_abs_inverses abs_inverses set_maps)
- pre_set_defss mss (unflat mss (1 upto n)) kkss)
+ mk_induct_discharge_prem_tac ctxt nn n pre_abs_inverses abs_inverses set_maps)
+ pre_set_defss mss (unflat mss (1 upto n)) kksss)
end;
-fun mk_coinduct_same_ctr_tac ctxt rel_eqs pre_rel_def fp_abs_inverse abs_inverse dtor_ctor ctr_def
- discs sels =
+fun mk_coinduct_same_ctr_tac ctxt rel_eqs pre_rel_def pre_abs_inverse abs_inverse dtor_ctor ctr_def
+ discs sels extra_unfolds =
hyp_subst_tac ctxt THEN'
CONVERSION (hhf_concl_conv
(Conv.top_conv (K (Conv.try_conv (Conv.rewr_conv ctr_def))) ctxt) ctxt) THEN'
SELECT_GOAL (unfold_thms_tac ctxt (pre_rel_def :: dtor_ctor :: sels)) THEN'
- SELECT_GOAL (unfold_thms_tac ctxt (pre_rel_def :: fp_abs_inverse :: abs_inverse :: dtor_ctor ::
- sels @ sumprod_thms_rel @ @{thms o_apply vimage2p_def})) THEN'
+ SELECT_GOAL (unfold_thms_tac ctxt (pre_rel_def :: pre_abs_inverse :: abs_inverse :: dtor_ctor ::
+ sels @ sumprod_thms_rel @ extra_unfolds @ @{thms o_apply vimage2p_def})) THEN'
(assume_tac ctxt ORELSE' REPEAT o etac ctxt conjE THEN'
full_simp_tac (ss_only (no_refl discs @ rel_eqs @ more_simp_thms) ctxt) THEN'
REPEAT o etac ctxt conjE THEN_MAYBE' REPEAT o hyp_subst_tac ctxt THEN'
@@ -368,8 +373,8 @@
full_simp_tac (ss_only (refl :: no_refl discs'' @ basic_simp_thms) ctxt)
end;
-fun mk_coinduct_discharge_prem_tac ctxt rel_eqs' nn kk n pre_rel_def fp_abs_inverse abs_inverse
- dtor_ctor exhaust ctr_defs discss selss =
+fun mk_coinduct_discharge_prem_tac ctxt extra_unfolds rel_eqs' nn kk n pre_rel_def pre_abs_inverse
+ abs_inverse dtor_ctor exhaust ctr_defs discss selss =
let val ks = 1 upto n in
EVERY' ([rtac ctxt allI, rtac ctxt allI, rtac ctxt impI,
select_prem_tac ctxt nn (dtac ctxt meta_spec) kk, dtac ctxt meta_spec, dtac ctxt meta_mp,
@@ -379,18 +384,18 @@
EVERY' ([rtac ctxt exhaust, K (co_induct_inst_as_projs_tac ctxt 1)] @
map2 (fn k' => fn discs' =>
if k' = k then
- mk_coinduct_same_ctr_tac ctxt rel_eqs' pre_rel_def fp_abs_inverse abs_inverse
- dtor_ctor ctr_def discs sels
+ mk_coinduct_same_ctr_tac ctxt rel_eqs' pre_rel_def pre_abs_inverse abs_inverse
+ dtor_ctor ctr_def discs sels extra_unfolds
else
mk_coinduct_distinct_ctrs_tac ctxt discs discs') ks discss)) ks ctr_defs discss selss)
end;
-fun mk_coinduct_tac ctxt rel_eqs' nn ns dtor_coinduct' pre_rel_defs fp_abs_inverses abs_inverses
+fun mk_coinduct_tac ctxt rel_eqs' nn ns dtor_coinduct' pre_rel_defs pre_abs_inverses abs_inverses
dtor_ctors exhausts ctr_defss discsss selsss =
HEADGOAL (rtac ctxt dtor_coinduct' THEN'
- EVERY' (@{map 10} (mk_coinduct_discharge_prem_tac ctxt rel_eqs' nn)
- (1 upto nn) ns pre_rel_defs fp_abs_inverses abs_inverses dtor_ctors exhausts ctr_defss discsss
- selsss));
+ EVERY' (@{map 10} (mk_coinduct_discharge_prem_tac ctxt [] rel_eqs' nn)
+ (1 upto nn) ns pre_rel_defs pre_abs_inverses abs_inverses dtor_ctors exhausts ctr_defss
+ discsss selsss));
fun mk_map_tac ctxt abs_inverses pre_map_def map_ctor live_nesting_map_id0s ctr_defs'
extra_unfolds =
--- a/src/HOL/Tools/BNF/bnf_fp_n2m_sugar.ML Wed Jan 04 14:26:08 2017 +0100
+++ b/src/HOL/Tools/BNF/bnf_fp_n2m_sugar.ML Wed Jan 04 14:26:19 2017 +0100
@@ -283,10 +283,10 @@
([induct], [inducts], rec_thmss, replicate nn [], replicate nn []))
||> (fn info => (SOME info, NONE))
else
- derive_coinduct_corecs_thms_for_types pre_bnfs (the corecs_args_types)
+ derive_coinduct_corecs_thms_for_types lthy pre_bnfs (the corecs_args_types)
xtor_co_induct dtor_injects dtor_ctors xtor_co_rec_thms live_nesting_bnfs fpTs Cs Xs
ctrXs_Tsss kss mss ns fp_abs_inverses abs_inverses
- (fn thm => thm RS @{thm vimage2p_refl}) ctr_defss ctr_sugars co_recs co_rec_defs lthy
+ (fn thm => thm RS @{thm vimage2p_refl}) ctr_defss ctr_sugars co_recs co_rec_defs
|> `(fn ((coinduct_thms_pairs, _), corec_thmss, corec_disc_thmss, _,
(corec_sel_thmsss, _)) =>
(map snd coinduct_thms_pairs, map fst coinduct_thms_pairs, corec_thmss,
--- a/src/HOL/Tools/BNF/bnf_fp_rec_sugar_util.ML Wed Jan 04 14:26:08 2017 +0100
+++ b/src/HOL/Tools/BNF/bnf_fp_rec_sugar_util.ML Wed Jan 04 14:26:19 2017 +0100
@@ -9,6 +9,44 @@
signature BNF_FP_REC_SUGAR_UTIL =
sig
val error_at: Proof.context -> term list -> string -> 'a
+ val warning_at: Proof.context -> term list -> string -> unit
+
+ val excess_equations: Proof.context -> term list -> 'a
+ val extra_variable_in_rhs: Proof.context -> term list -> term -> 'a
+ val ill_formed_corec_call: Proof.context -> term -> 'a
+ val ill_formed_equation_head: Proof.context -> term list -> 'a
+ val ill_formed_equation_lhs_rhs: Proof.context -> term list -> 'a
+ val ill_formed_equation: Proof.context -> term -> 'a
+ val ill_formed_formula: Proof.context -> term -> 'a
+ val ill_formed_rec_call: Proof.context -> term -> 'a
+ val inconstant_pattern_pos_for_fun: Proof.context -> term list -> string -> 'a
+ val invalid_map: Proof.context -> term list -> term -> 'a
+ val missing_args_to_fun_on_lhs: Proof.context -> term list -> 'a
+ val missing_equations_for_const: string -> 'a
+ val missing_equations_for_fun: string -> 'a
+ val missing_pattern: Proof.context -> term list -> 'a
+ val more_than_one_nonvar_in_lhs: Proof.context -> term list -> 'a
+ val multiple_equations_for_ctr: Proof.context -> term list -> 'a
+ val nonprimitive_corec: Proof.context -> term list -> 'a
+ val nonprimitive_pattern_in_lhs: Proof.context -> term list -> 'a
+ val not_codatatype: Proof.context -> typ -> 'a
+ val not_datatype: Proof.context -> typ -> 'a
+ val not_constructor_in_pattern: Proof.context -> term list -> term -> 'a
+ val not_constructor_in_rhs: Proof.context -> term list -> term -> 'a
+ val rec_call_not_apply_to_ctr_arg: Proof.context -> term list -> term -> 'a
+ val partially_applied_ctr_in_pattern: Proof.context -> term list -> 'a
+ val partially_applied_ctr_in_rhs: Proof.context -> term list -> 'a
+ val too_few_args_in_rec_call: Proof.context -> term list -> term -> 'a
+ val unexpected_rec_call_in: Proof.context -> term list -> term -> 'a
+ val unexpected_corec_call_in: Proof.context -> term list -> term -> 'a
+ val unsupported_case_around_corec_call: Proof.context -> term list -> term -> 'a
+
+ val no_equation_for_ctr_warning: Proof.context -> term list -> term -> unit
+
+ val check_all_fun_arg_frees: Proof.context -> term list -> term list -> unit
+ val check_duplicate_const_names: binding list -> unit
+ val check_duplicate_variables_in_lhs: Proof.context -> term list -> term list -> unit
+ val check_top_sort: Proof.context -> binding -> typ -> unit
datatype fp_kind = Least_FP | Greatest_FP
@@ -56,14 +94,116 @@
val mk_conjunctN: int -> int -> thm
val conj_dests: int -> thm -> thm list
+
+ val print_def_consts: bool -> (term * (string * thm)) list -> Proof.context -> unit
end;
structure BNF_FP_Rec_Sugar_Util : BNF_FP_REC_SUGAR_UTIL =
struct
-fun error_at ctxt ts str =
- error (str ^ (if null ts then ""
- else " at\n " ^ space_implode "\n " (map (quote o Syntax.string_of_term ctxt) ts)));
+fun error_at ctxt ats str =
+ error (str ^ (if null ats then ""
+ else " at\n" ^ cat_lines (map (prefix " " o Syntax.string_of_term ctxt) ats)));
+fun warning_at ctxt ats str =
+ warning (str ^ (if null ats then ""
+ else " at\n" ^ cat_lines (map (prefix " " o Syntax.string_of_term ctxt) ats)));
+
+fun excess_equations ctxt ats =
+ error ("Excess equation(s):\n" ^
+ cat_lines (map (prefix " " o Syntax.string_of_term ctxt) ats));
+fun extra_variable_in_rhs ctxt ats var =
+ error_at ctxt ats ("Extra variable " ^ quote (Syntax.string_of_term ctxt var) ^
+ " in right-hand side");
+fun ill_formed_corec_call ctxt t =
+ error ("Ill-formed corecursive call " ^ quote (Syntax.string_of_term ctxt t));
+fun ill_formed_equation_head ctxt ats =
+ error_at ctxt ats "Ill-formed function equation (expected function name on left-hand side)";
+fun ill_formed_equation_lhs_rhs ctxt ats =
+ error_at ctxt ats "Ill-formed equation (expected \"lhs = rhs\")";
+fun ill_formed_equation ctxt t =
+ error_at ctxt [] ("Ill-formed equation:\n " ^ Syntax.string_of_term ctxt t);
+fun ill_formed_formula ctxt t =
+ error_at ctxt [] ("Ill-formed formula:\n " ^ Syntax.string_of_term ctxt t);
+fun ill_formed_rec_call ctxt t =
+ error ("Ill-formed recursive call: " ^ quote (Syntax.string_of_term ctxt t));
+fun inconstant_pattern_pos_for_fun ctxt ats fun_name =
+ error_at ctxt ats ("Inconstant constructor pattern position for function " ^ quote fun_name);
+fun invalid_map ctxt ats t =
+ error_at ctxt ats ("Invalid map function in " ^ quote (Syntax.string_of_term ctxt t));
+fun missing_args_to_fun_on_lhs ctxt ats =
+ error_at ctxt ats "Expected more arguments to function on left-hand side";
+fun missing_equations_for_const fun_name =
+ error ("Missing equations for constant " ^ quote fun_name);
+fun missing_equations_for_fun fun_name =
+ error ("Missing equations for function " ^ quote fun_name);
+fun missing_pattern ctxt ats =
+ error_at ctxt ats "Constructor pattern missing in left-hand side";
+fun more_than_one_nonvar_in_lhs ctxt ats =
+ error_at ctxt ats "More than one non-variable argument in left-hand side";
+fun multiple_equations_for_ctr ctxt ats =
+ error ("Multiple equations for constructor:\n" ^
+ cat_lines (map (prefix " " o Syntax.string_of_term ctxt) ats));
+fun nonprimitive_corec ctxt ats =
+ error_at ctxt ats "Nonprimitive corecursive specification";
+fun nonprimitive_pattern_in_lhs ctxt ats =
+ error_at ctxt ats "Nonprimitive pattern in left-hand side";
+fun not_codatatype ctxt T =
+ error ("Not a codatatype: " ^ Syntax.string_of_typ ctxt T);
+fun not_datatype ctxt T =
+ error ("Not a datatype: " ^ Syntax.string_of_typ ctxt T);
+fun not_constructor_in_pattern ctxt ats t =
+ error_at ctxt ats ("Not a constructor " ^ quote (Syntax.string_of_term ctxt t) ^
+ " in pattern");
+fun not_constructor_in_rhs ctxt ats t =
+ error_at ctxt ats ("Not a constructor " ^ quote (Syntax.string_of_term ctxt t) ^
+ " in right-hand side");
+fun rec_call_not_apply_to_ctr_arg ctxt ats t =
+ error_at ctxt ats ("Recursive call not directly applied to constructor argument in " ^
+ quote (Syntax.string_of_term ctxt t));
+fun partially_applied_ctr_in_pattern ctxt ats =
+ error_at ctxt ats "Partially applied constructor in pattern";
+fun partially_applied_ctr_in_rhs ctxt ats =
+ error_at ctxt ats "Partially applied constructor in right-hand side";
+fun too_few_args_in_rec_call ctxt ats t =
+ error_at ctxt ats ("Too few arguments in recursive call " ^ quote (Syntax.string_of_term ctxt t));
+fun unexpected_rec_call_in ctxt ats t =
+ error_at ctxt ats ("Unexpected recursive call in " ^ quote (Syntax.string_of_term ctxt t));
+fun unexpected_corec_call_in ctxt ats t =
+ error_at ctxt ats ("Unexpected corecursive call in " ^ quote (Syntax.string_of_term ctxt t));
+fun unsupported_case_around_corec_call ctxt ats t =
+ error_at ctxt ats ("Unsupported corecursive call under case expression " ^
+ quote (Syntax.string_of_term ctxt t) ^ "\n(Define " ^
+ quote (Syntax.string_of_typ ctxt (domain_type (fastype_of t))) ^
+ " with discriminators and selectors to circumvent this limitation)");
+
+fun no_equation_for_ctr_warning ctxt ats ctr =
+ warning_at ctxt ats ("No equation for constructor " ^ quote (Syntax.string_of_term ctxt ctr));
+
+fun check_all_fun_arg_frees ctxt ats fun_args =
+ (case find_first (not o is_Free) fun_args of
+ SOME t => error_at ctxt ats ("Non-variable function argument on left-hand side " ^
+ quote (Syntax.string_of_term ctxt t))
+ | NONE =>
+ (case find_first (Variable.is_fixed ctxt o fst o dest_Free) fun_args of
+ SOME t => error_at ctxt ats ("Function argument " ^
+ quote (Syntax.string_of_term ctxt t) ^ " is fixed in context")
+ | NONE => ()));
+
+fun check_duplicate_const_names bs =
+ let val dups = duplicates (op =) (map Binding.name_of bs) in
+ ignore (null dups orelse error ("Duplicate constant name " ^ quote (hd dups)))
+ end;
+
+fun check_duplicate_variables_in_lhs ctxt ats vars =
+ let val dups = duplicates (op aconv) vars in
+ ignore (null dups orelse
+ error_at ctxt ats ("Duplicable variable " ^ quote (Syntax.string_of_term ctxt (hd dups)) ^
+ " in left-hand side"))
+ end;
+
+fun check_top_sort ctxt b T =
+ ignore (Sign.of_sort (Proof_Context.theory_of ctxt) (T, @{sort type}) orelse
+ error ("Type of " ^ Binding.print b ^ " contains top sort"));
datatype fp_kind = Least_FP | Greatest_FP;
@@ -174,4 +314,8 @@
fun conj_dests n thm = map (fn k => thm RS mk_conjunctN n k) (1 upto n);
+fun print_def_consts int defs ctxt =
+ Proof_Display.print_consts int (Position.thread_data ()) ctxt (K false)
+ (map_filter (try (dest_Free o fst)) defs);
+
end;
--- a/src/HOL/Tools/BNF/bnf_gfp_grec.ML Wed Jan 04 14:26:08 2017 +0100
+++ b/src/HOL/Tools/BNF/bnf_gfp_grec.ML Wed Jan 04 14:26:19 2017 +0100
@@ -261,7 +261,7 @@
val (([free], [def_free], [simps_free]), (lthy, lthy_old)) = lthy
|> Local_Theory.open_target |> snd
|> Local_Theory.map_background_naming (mk_internal true lthy Name_Space.concealed) (*TODO check*)
- |> primrec [] [(b, NONE, NoSyn)] (map (fn eq => ((Binding.empty_atts, eq), [], [])) eqs)
+ |> primrec false [] [(b, NONE, NoSyn)] (map (fn eq => ((Binding.empty_atts, eq), [], [])) eqs)
||> `Local_Theory.close_target;
val phi = Proof_Context.export_morphism lthy_old lthy;
@@ -1005,8 +1005,8 @@
val simple_Ts = map fst simple_T_mapfs;
- val t_map = build_map ctxt simple_Ts build_simple (apply2 (range_type o fastype_of) (t, u));
- val u_map = build_map ctxt simple_Ts build_simple (apply2 (domain_type o fastype_of) (t, u));
+ val t_map = build_map ctxt simple_Ts [] build_simple (apply2 (range_type o fastype_of) (t, u));
+ val u_map = build_map ctxt simple_Ts [] build_simple (apply2 (domain_type o fastype_of) (t, u));
in
mk_Trueprop_eq (HOLogic.mk_comp (u, u_map), HOLogic.mk_comp (t_map, t))
end;
--- a/src/HOL/Tools/BNF/bnf_gfp_grec_sugar.ML Wed Jan 04 14:26:08 2017 +0100
+++ b/src/HOL/Tools/BNF/bnf_gfp_grec_sugar.ML Wed Jan 04 14:26:19 2017 +0100
@@ -23,10 +23,10 @@
val derive_unique: Proof.context -> morphism -> term -> BNF_GFP_Grec.corec_info -> string ->
thm -> thm
- val corec_cmd: corec_option list -> (binding * string option * mixfix) list * string ->
+ val corec_cmd: bool -> corec_option list -> (binding * string option * mixfix) list * string ->
local_theory -> local_theory
- val corecursive_cmd: corec_option list -> (binding * string option * mixfix) list * string ->
- local_theory -> Proof.state
+ val corecursive_cmd: bool -> corec_option list ->
+ (binding * string option * mixfix) list * string -> local_theory -> Proof.state
val friend_of_corec_cmd: (string * string option) * string -> local_theory -> Proof.state
val coinduction_upto_cmd: string * string -> local_theory -> local_theory
end;
@@ -87,14 +87,6 @@
#> Raw_Simplifier.rewrite_term thy @{thms BNF_Composition.id_bnf_def} []
end;
-fun unexpected_corec_call ctxt eqns t =
- error_at ctxt eqns ("Unexpected corecursive call in " ^ quote (Syntax.string_of_term ctxt t));
-fun unsupported_case_around_corec_call ctxt eqns t =
- error_at ctxt eqns ("Unsupported corecursive call under case expression " ^
- quote (Syntax.string_of_term ctxt t) ^ "\n(Define " ^
- quote (Syntax.string_of_typ ctxt (domain_type (fastype_of t))) ^
- " with discriminators and selectors to circumvent this limitation.)");
-
datatype corec_option =
Plugins_Option of Proof.context -> Plugin_Name.filter |
Friend_Option |
@@ -587,7 +579,7 @@
||>> mk_Frees "y" xy_Ts;
fun mk_prem xy_T x y =
- build_rel [] ctxt [fpT] (fn (T, _) => if T = fpT then Rcong else HOLogic.eq_const T)
+ build_rel [] ctxt [fpT] [] (fn (T, _) => if T = fpT then Rcong else HOLogic.eq_const T)
(xy_T, xy_T) $ x $ y;
val prems = @{map 3} mk_prem xy_Ts xs ys;
@@ -685,9 +677,9 @@
fun mk_applied_cong arg =
enforce_type ctxt domain_type (fastype_of arg) cong0 $ arg;
- val thm = derive_coinduct_thms_for_types false mk_applied_cong [pre_bnf] dtor_coinduct
+ val thm = derive_coinduct_thms_for_types ctxt false mk_applied_cong [pre_bnf] dtor_coinduct
dtor_ctors live_nesting_bnfs [fpT] [substXAT X] [map (map substXAT) ctrXs_Tss] [n]
- [abs_inverse] [abs_inverse] I [ctr_defs] [morph_ctr_sugar phi ctr_sugar0] ctxt
+ [abs_inverse] [abs_inverse] I [ctr_defs] [morph_ctr_sugar phi ctr_sugar0]
|> map snd |> the_single;
val (attrs, _) = mk_coinduct_attrs [fpT] [ctrs0] [discs0] [ms];
in
@@ -713,7 +705,7 @@
|> explore {bound_Us = T :: bound_Us, bound_Ts = T :: bound_Ts, U = U, T = T}
|> (fn t => Abs (Name.uu, T, t));
in
- betapply (build_map lthy [] build_simple (T, U), t)
+ betapply (build_map lthy [] [] build_simple (T, U), t)
end;
fun add_boundvar t = betapply (incr_boundvars 1 t, Bound 0);
@@ -754,43 +746,24 @@
fun check_fun_name () =
null fun_frees orelse member (op aconv) fun_frees fun_t orelse
- error (quote (Syntax.string_of_term ctxt fun_t) ^
- " is not the function currently being defined");
-
- fun check_args_are_vars () =
- let
- fun is_ok_Free_or_Var (Free (s, _)) = not (String.isSuffix inner_fp_suffix s)
- | is_ok_Free_or_Var (Var _) = true
- | is_ok_Free_or_Var _ = false;
-
- fun is_valid arg =
- (is_ok_Free_or_Var arg andalso not (member (op aconv) fun_frees arg)) orelse
- error ("Argument " ^ quote (Syntax.string_of_term ctxt arg) ^ " is not free");
- in
- forall is_valid arg_ts
- end;
-
- fun check_no_duplicate_arg () =
- (case duplicates (op =) arg_ts of
- [] => ()
- | arg :: _ => error ("Repeated argument: " ^ quote (Syntax.string_of_term ctxt arg)));
+ ill_formed_equation_head ctxt [] fun_t;
fun check_no_other_frees () =
(case Term.add_frees rhs [] |> map Free |> subtract (op =) (fun_frees @ arg_ts)
- |> filter_out (Variable.is_fixed ctxt o fst o dest_Free) of
- [] => ()
- | Free (s, _) :: _ => error ("Extra variable on right-hand side: " ^ quote s));
+ |> find_first (not o Variable.is_fixed ctxt o fst o dest_Free) of
+ NONE => ()
+ | SOME t => extra_variable_in_rhs ctxt [] t);
in
- check_no_duplicate_arg ();
+ check_duplicate_variables_in_lhs ctxt [] arg_ts;
check_fun_name ();
- check_args_are_vars ();
+ check_all_fun_arg_frees ctxt [] (filter_out is_Var arg_ts);
check_no_other_frees ()
end;
fun parse_corec_equation ctxt fun_frees eq =
let
val (lhs, rhs) = HOLogic.dest_eq (HOLogic.dest_Trueprop (drop_all eq))
- handle TERM _ => error "Expected HOL equation";
+ handle TERM _ => ill_formed_equation_lhs_rhs ctxt [eq];
val _ = check_corec_equation ctxt fun_frees (lhs, rhs);
@@ -1047,7 +1020,7 @@
val param2_T = Type_Infer.paramify_vars param1_T;
val deadfixed_T =
- build_map lthy [] (mk_undefined o op -->) (apply2 tupleT_of_funT (param1_T, param2_T))
+ build_map lthy [] [] (mk_undefined o op -->) (apply2 tupleT_of_funT (param1_T, param2_T))
|> singleton (Type_Infer_Context.infer_types lthy)
|> singleton (Type_Infer.fixate lthy false)
|> type_of
@@ -1159,24 +1132,34 @@
|> `(curry fastype_of1 bound_Ts)
|>> build_params bound_Us bound_Ts
|-> explore;
- val Us = map (fpT_to ssig_T) (snd (dest_Type (fastype_of1 (bound_Us, mapped_arg'))));
- val temporary_map = map_tm
- |> mk_map n Us Ts;
- val map_fn_Ts = fastype_of #> strip_fun_type #> fst;
- val binder_Uss = map_fn_Ts temporary_map
- |> map (map (fpT_to ssig_T) o binder_types);
- val fun_paramss = map_fn_Ts (head_of func)
- |> map (build_params bound_Us bound_Ts);
- val fs' = fs |> @{map 4} explore_fun binder_Uss (replicate n explore) fun_paramss;
- val SOME bnf = bnf_of lthy T_name;
- val Type (_, bnf_Ts) = T_of_bnf bnf;
- val typ_alist =
- lives_of_bnf bnf ~~ map (curry fastype_of1 bound_Us #> range_type) fs';
- val Us' = map2 the_default Us (map (AList.lookup (op =) typ_alist) bnf_Ts);
- val map_tm' = map_tm |> mk_map n Us Us';
in
- build_function_after_encapsulation func (list_comb (map_tm', fs')) params [mapped_arg]
- [mapped_arg']
+ (case fastype_of1 (bound_Us, mapped_arg') of
+ Type (U_name, Us0) =>
+ if U_name = T_name then
+ let
+ val Us = map (fpT_to ssig_T) Us0;
+ val temporary_map = map_tm
+ |> mk_map n Us Ts;
+ val map_fn_Ts = fastype_of #> strip_fun_type #> fst;
+ val binder_Uss = map_fn_Ts temporary_map
+ |> map (map (fpT_to ssig_T) o binder_types);
+ val fun_paramss = map_fn_Ts (head_of func)
+ |> map (build_params bound_Us bound_Ts);
+ val fs' = fs
+ |> @{map 4} explore_fun binder_Uss (replicate n explore) fun_paramss;
+ val SOME bnf = bnf_of lthy T_name;
+ val Type (_, bnf_Ts) = T_of_bnf bnf;
+ val typ_alist =
+ lives_of_bnf bnf ~~ map (curry fastype_of1 bound_Us #> range_type) fs';
+ val Us' = map2 the_default Us (map (AList.lookup (op =) typ_alist) bnf_Ts);
+ val map_tm' = map_tm |> mk_map n Us Us';
+ in
+ build_function_after_encapsulation func (list_comb (map_tm', fs')) params
+ [mapped_arg] [mapped_arg']
+ end
+ else
+ explore params t
+ | _ => explore params t)
end
| NONE => explore params t)
| massage_map explore params t = explore params t;
@@ -1526,10 +1509,7 @@
| NONE => explore params t)
| _ => explore params t)
and explore_cond params t =
- if has_self_call t then
- error ("Unallowed corecursive call in condition " ^ quote (Syntax.string_of_term lthy t))
- else
- explore_inner params t
+ if has_self_call t then unexpected_rec_call_in lthy [] t else explore_inner params t
and explore_inner params t =
massage_rho explore_inner_general params t
and explore_inner_general (params as {bound_Us, bound_Ts, T, ...}) t =
@@ -1550,7 +1530,7 @@
disc' $ arg'
else
error ("Discriminator " ^ quote (Syntax.string_of_term lthy fun_t) ^
- " cannot be applied to non-lhs argument " ^
+ " cannot be applied to non-variable " ^
quote (Syntax.string_of_term lthy (hd arg_ts)))
end
| _ =>
@@ -1566,7 +1546,7 @@
build_function_after_encapsulation fun_t sel' params arg_ts arg_ts'
else
error ("Selector " ^ quote (Syntax.string_of_term lthy fun_t) ^
- " cannot be applied to non-lhs argument " ^
+ " cannot be applied to non-variable " ^
quote (Syntax.string_of_term lthy (hd arg_ts)))
end
| NONE =>
@@ -1605,8 +1585,7 @@
if is_some ctr_opt then
rebuild_function_after_exploration fun_t (the ctr_opt) explore_inner params arg_ts
else
- error ("Constructor expected on right-hand side, " ^
- quote (Syntax.string_of_term lthy fun_t) ^ " found instead")
+ not_constructor_in_rhs lthy [] fun_t
end;
val rho_rhs = rhs
@@ -1638,11 +1617,9 @@
HOLogic.mk_Trueprop (rho_rel $ rho_rhs $ rho_rhsZ)
end;
-fun extract_rho_return_transfer_goals fun_b version dead_pre_bnf dead_k_bnf Y Z preT fun_T k_T
- ssig_T ssig_fp_sugar friend_parse_info fun_t parsed_eq lthy =
+fun extract_rho_return_transfer_goals fun_b version dead_pre_bnf dead_k_bnf Y Z preT k_T ssig_T
+ ssig_fp_sugar friend_parse_info fun_t parsed_eq lthy =
let
- val Type (fpT_name, _) = body_type fun_T;
-
fun mk_rel T bnf =
let
val ZT = Tsubst Y Z T;
@@ -1724,7 +1701,7 @@
fun massage_let_if_case_corec explore {bound_Ts, U, T, ...} t =
massage_let_if_case ctxt has_self_call (fn bound_Ts => explore (build_params bound_Ts U T))
- (K (unexpected_corec_call ctxt [t])) (K (unsupported_case_around_corec_call ctxt [t]))
+ (K (unexpected_corec_call_in ctxt [t])) (K (unsupported_case_around_corec_call ctxt [t]))
bound_Ts t;
val massage_map_let_if_case =
@@ -1985,7 +1962,7 @@
else
((false, extras), lthy));
-fun prepare_corec_ursive_cmd long_cmd opts (raw_fixes, raw_eq) lthy =
+fun prepare_corec_ursive_cmd int long_cmd opts (raw_fixes, raw_eq) lthy =
let
val _ = can the_single raw_fixes orelse
error "Mutually corecursive functions not supported";
@@ -1994,8 +1971,7 @@
val ([((b, fun_T), mx)], [(_, eq)]) =
fst (Specification.read_multi_specs raw_fixes [((Binding.empty_atts, raw_eq), [], [])] lthy);
- val _ = Sign.of_sort (Proof_Context.theory_of lthy) (fun_T, @{sort type}) orelse
- error ("Type of " ^ Binding.print b ^ " contains top sort");
+ val _ = check_top_sort lthy b fun_T;
val (arg_Ts, res_T) = strip_type fun_T;
val fpT_name = (case res_T of Type (s, _) => s | _ => not_codatatype lthy res_T);
@@ -2020,8 +1996,8 @@
val parsed_eq' = parsed_eq ||> subst_atomic [(fun_free, fun_t)];
in
lthy
- |> extract_rho_return_transfer_goals b version dead_pre_bnf dead_k_bnf Y Z preT fun_T k_T
- ssig_T ssig_fp_sugar friend_parse_info fun_t parsed_eq'
+ |> extract_rho_return_transfer_goals b version dead_pre_bnf dead_k_bnf Y Z preT k_T ssig_T
+ ssig_fp_sugar friend_parse_info fun_t parsed_eq'
|>> pair prepared
end;
@@ -2080,6 +2056,7 @@
val ((fun_t0, (_, fun_def0)), (lthy, lthy_old)) = lthy
|> Local_Theory.open_target |> snd
|> Local_Theory.define def
+ |> tap (fn (def, lthy) => print_def_consts int [def] lthy)
||> `Local_Theory.close_target;
val parsed_eq = parse_corec_equation lthy [fun_free] eq;
@@ -2207,11 +2184,11 @@
(inner_fp_triple, termin_goals), (const_transfers, const_transfer_goals'))), lthy)
end;
-fun corec_cmd opts (raw_fixes, raw_eq) lthy =
+fun corec_cmd int opts (raw_fixes, raw_eq) lthy =
let
val ((def_fun, (_, (inner_fp_triple, termin_goals), (const_transfers, const_transfer_goals))),
lthy) =
- prepare_corec_ursive_cmd false opts (raw_fixes, raw_eq) lthy;
+ prepare_corec_ursive_cmd int false opts (raw_fixes, raw_eq) lthy;
in
if not (null termin_goals) then
error ("Termination prover failed (try " ^ quote (#1 @{command_keyword corecursive}) ^
@@ -2223,11 +2200,11 @@
def_fun inner_fp_triple const_transfers [] lthy
end;
-fun corecursive_cmd opts (raw_fixes, raw_eq) lthy =
+fun corecursive_cmd int opts (raw_fixes, raw_eq) lthy =
let
val ((def_fun, (([Type (fpT_name, _)], prepareds, rho_datas, rho_transfer_goals),
(inner_fp_triple, termin_goals), (const_transfers, const_transfer_goals))), lthy) =
- prepare_corec_ursive_cmd true opts (raw_fixes, raw_eq) lthy;
+ prepare_corec_ursive_cmd int true opts (raw_fixes, raw_eq) lthy;
val (rho_transfer_goals', unprime_rho_transfer_and_folds) =
@{map 3} (fn (_, _, _, _, _, _, _, _, _, _, _, _, _) => fn (_, rho_def) =>
@@ -2266,7 +2243,7 @@
val fun_T =
(case code_goal of
@{const Trueprop} $ (Const (@{const_name HOL.eq}, _) $ t $ _) => fastype_of (head_of t)
- | _ => error "Expected HOL equation");
+ | _ => ill_formed_equation_lhs_rhs lthy [code_goal]);
val fun_t = Const (fun_name, fun_T);
val (arg_Ts, res_T as Type (fpT_name, _)) = strip_type fun_T;
@@ -2283,8 +2260,8 @@
val parsed_eq = parse_corec_equation lthy [] code_goal;
val (((rho, rho_def), (const_transfer_goals, rho_transfer_goal)), lthy) =
- extract_rho_return_transfer_goals fun_b version dead_pre_bnf dead_k_bnf Y Z preT fun_T k_T
- ssig_T ssig_fp_sugar friend_parse_info fun_t parsed_eq lthy;
+ extract_rho_return_transfer_goals fun_b version dead_pre_bnf dead_k_bnf Y Z preT k_T ssig_T
+ ssig_fp_sugar friend_parse_info fun_t parsed_eq lthy;
fun register_friend_extra_and_note_thms code_goal code_thm const_transfers k_T friend_info
lthy =
@@ -2393,13 +2370,13 @@
"define nonprimitive corecursive functions"
((Scan.optional (@{keyword "("} |-- Parse.!!! (Parse.list1 corec_option_parser)
--| @{keyword ")"}) []) -- (Parse.vars --| Parse.where_ -- Parse.prop)
- >> uncurry corec_cmd);
+ >> uncurry (corec_cmd true));
val _ = Outer_Syntax.local_theory_to_proof @{command_keyword corecursive}
"define nonprimitive corecursive functions"
((Scan.optional (@{keyword "("} |-- Parse.!!! (Parse.list1 corec_option_parser)
--| @{keyword ")"}) []) -- (Parse.vars --| Parse.where_ -- Parse.prop)
- >> uncurry corecursive_cmd);
+ >> uncurry (corecursive_cmd true));
val _ = Outer_Syntax.local_theory_to_proof @{command_keyword friend_of_corec}
"register a function as a legal context for nonprimitive corecursion"
--- a/src/HOL/Tools/BNF/bnf_gfp_grec_sugar_util.ML Wed Jan 04 14:26:08 2017 +0100
+++ b/src/HOL/Tools/BNF/bnf_gfp_grec_sugar_util.ML Wed Jan 04 14:26:19 2017 +0100
@@ -49,9 +49,6 @@
val mk_case_sumN_balanced = Balanced_Tree.make mk_case_sum;
-fun not_codatatype ctxt T =
- error ("Not a codatatype: " ^ Syntax.string_of_typ ctxt T);
-
fun generalize_types max_j T U =
let
val vars = Unsynchronized.ref [];
@@ -156,7 +153,8 @@
Variable.add_free_names ctxt goal []
|> (fn vars => Goal.prove_sorry ctxt vars [] goal (fn {context = ctxt, prems = _} =>
unfold_thms_tac ctxt [rel_def] THEN
- HEADGOAL (rtac ctxt (@{thm Abs_transfer} OF [type_definition, type_definition]))))
+ HEADGOAL (rtac ctxt refl ORELSE'
+ rtac ctxt (@{thm Abs_transfer} OF [type_definition, type_definition]))))
end
end;
@@ -178,7 +176,7 @@
Variable.add_free_names ctxt goal []
|> (fn vars => Goal.prove_sorry ctxt vars [] goal (fn {context = ctxt, prems = _} =>
unfold_thms_tac ctxt [rel_def] THEN
- HEADGOAL (rtac ctxt @{thm vimage2p_rel_fun})))
+ HEADGOAL (rtac ctxt refl ORELSE' rtac ctxt @{thm vimage2p_rel_fun})))
end
end;
@@ -208,7 +206,7 @@
if op = TU then
t
else
- (case try (build_map ctxt [] (the o AList.lookup (op =) AB_fs)) TU of
+ (case try (build_map ctxt [] [] (the o AList.lookup (op =) AB_fs)) TU of
SOME mapx => mapx $ t
| NONE => raise UNNATURAL ());
@@ -355,7 +353,7 @@
fun mk_friend_spec () =
let
fun encapsulate_nested U T free =
- betapply (build_map ctxt [] (fn (T, _) =>
+ betapply (build_map ctxt [] [] (fn (T, _) =>
if T = domain_type VLeaf0_T then Abs (Name.uu, T, VLeaf0 $ Bound 0)
else Abs (Name.uu, T, Bound 0)) (T, U),
free);
--- a/src/HOL/Tools/BNF/bnf_gfp_rec_sugar.ML Wed Jan 04 14:26:08 2017 +0100
+++ b/src/HOL/Tools/BNF/bnf_gfp_rec_sugar.ML Wed Jan 04 14:26:19 2017 +0100
@@ -82,17 +82,17 @@
val gfp_rec_sugar_interpretation: string ->
(BNF_FP_Rec_Sugar_Util.fp_rec_sugar -> local_theory -> local_theory) -> theory -> theory
- val primcorec_ursive: bool -> corec_option list -> ((binding * typ) * mixfix) list ->
+ val primcorec_ursive: bool -> bool -> corec_option list -> ((binding * typ) * mixfix) list ->
((binding * Token.T list list) * term) list -> term option list -> Proof.context ->
(term * 'a list) list list * (thm list list -> local_theory -> local_theory) * local_theory
- val primcorec_ursive_cmd: bool -> corec_option list ->
+ val primcorec_ursive_cmd: bool -> bool -> corec_option list ->
(binding * string option * mixfix) list * ((Attrib.binding * string) * string option) list ->
Proof.context ->
(term * 'a list) list list * (thm list list -> local_theory -> local_theory) * local_theory
- val primcorecursive_cmd: corec_option list ->
+ val primcorecursive_cmd: bool -> corec_option list ->
(binding * string option * mixfix) list * ((Attrib.binding * string) * string option) list ->
Proof.context -> Proof.state
- val primcorec_cmd: corec_option list ->
+ val primcorec_cmd: bool -> corec_option list ->
(binding * string option * mixfix) list * ((Attrib.binding * string) * string option) list ->
local_theory -> local_theory
end;
@@ -121,21 +121,6 @@
val nitpicksimp_attrs = @{attributes [nitpick_simp]};
val simp_attrs = @{attributes [simp]};
-fun extra_variable ctxt ts var =
- error_at ctxt ts ("Extra variable " ^ quote (Syntax.string_of_term ctxt var));
-fun not_codatatype ctxt T =
- error ("Not a codatatype: " ^ Syntax.string_of_typ ctxt T);
-fun ill_formed_corec_call ctxt t =
- error ("Ill-formed corecursive call " ^ quote (Syntax.string_of_term ctxt t));
-fun invalid_map ctxt t =
- error_at ctxt [t] "Invalid map function";
-fun nonprimitive_corec ctxt eqns =
- error_at ctxt eqns "Nonprimitive corecursive specification";
-fun unexpected_corec_call ctxt eqns t =
- error_at ctxt eqns ("Unexpected corecursive call in " ^ quote (Syntax.string_of_term ctxt t));
-fun unsupported_case_around_corec_call ctxt eqns t =
- error_at ctxt eqns ("Unsupported corecursive call under case expression " ^
- quote (Syntax.string_of_term ctxt t) ^ " for datatype with no discriminators and selectors");
fun use_primcorecursive () =
error ("\"auto\" failed (try " ^ quote (#1 @{command_keyword primcorecursive}) ^ " instead of " ^
quote (#1 @{command_keyword primcorec}) ^ ")");
@@ -352,12 +337,12 @@
end;
fun massage_let_if_case_corec ctxt has_call massage_leaf bound_Ts t0 =
- massage_let_if_case ctxt has_call massage_leaf (K (unexpected_corec_call ctxt [t0]))
+ massage_let_if_case ctxt has_call massage_leaf (K (unexpected_corec_call_in ctxt [t0]))
(K (unsupported_case_around_corec_call ctxt [t0])) bound_Ts t0;
fun massage_nested_corec_call ctxt has_call massage_call massage_noncall bound_Ts U T t0 =
let
- fun check_no_call t = if has_call t then unexpected_corec_call ctxt [t0] t else ();
+ fun check_no_call t = if has_call t then unexpected_corec_call_in ctxt [t0] t else ();
fun massage_mutual_call bound_Ts (Type (@{type_name fun}, [_, U2]))
(Type (@{type_name fun}, [T1, T2])) t =
@@ -658,23 +643,12 @@
fun check_extra_frees ctxt frees names t =
let val bads = add_extra_frees ctxt frees names t [] in
- null bads orelse extra_variable ctxt [t] (hd bads)
+ null bads orelse extra_variable_in_rhs ctxt [t] (hd bads)
end;
fun check_fun_args ctxt eqn fun_args =
- let
- val dups = duplicates (op =) fun_args;
- val _ = null dups orelse error_at ctxt [eqn]
- ("Duplicate variable " ^ quote (Syntax.string_of_term ctxt (hd dups)));
-
- val _ = forall is_Free fun_args orelse
- error_at ctxt [eqn] ("Non-variable function argument on left-hand side " ^
- quote (Syntax.string_of_term ctxt (the (find_first (not o is_Free) fun_args))));
-
- val fixeds = filter (Variable.is_fixed ctxt o fst o dest_Free) fun_args;
- val _ = null fixeds orelse error_at ctxt [eqn] ("Function argument " ^
- quote (Syntax.string_of_term ctxt (hd fixeds)) ^ " is fixed in context");
- in () end;
+ (check_duplicate_variables_in_lhs ctxt [eqn] fun_args;
+ check_all_fun_arg_frees ctxt [eqn] fun_args);
fun dissect_coeqn_disc ctxt fun_names sequentials
(basic_ctr_specss : basic_corec_ctr_spec list list) eqn_pos ctr_rhs_opt code_rhs_opt prems0
@@ -695,7 +669,7 @@
val _ = check_fun_args ctxt concl fun_args;
val bads = filter (Term.exists_subterm (is_free_in fun_names)) prems0;
- val _ = null bads orelse error_at ctxt bads "Corecursive call(s) in condition(s)";
+ val _ = null bads orelse unexpected_rec_call_in ctxt [] (hd bads);
val (sequential, basic_ctr_specs) =
the (AList.lookup (op =) (fun_names ~~ (sequentials ~~ basic_ctr_specss)) fun_name);
@@ -752,7 +726,7 @@
ctr_rhs_opt code_rhs_opt eqn0 of_spec_opt eqn =
let
val (lhs, rhs) = HOLogic.dest_eq eqn
- handle TERM _ => error_at ctxt [eqn] "Ill-formed equation (expected \"lhs = rhs\")";
+ handle TERM _ => ill_formed_equation_lhs_rhs ctxt [eqn];
val sel = head_of lhs;
val ((fun_name, fun_T), fun_args) = dest_comb lhs |> snd |> strip_comb |> apfst dest_Free
@@ -787,7 +761,7 @@
val basic_ctr_specs = the (AList.lookup (op =) (fun_names ~~ basic_ctr_specss) fun_name);
val (ctr, ctr_args) = strip_comb (unfold_lets_splits rhs);
val {disc, sels, ...} = the (find_first (curry (op =) ctr o #ctr) basic_ctr_specs)
- handle Option.Option => error_at ctxt [ctr] "Not a constructor";
+ handle Option.Option => not_constructor_in_rhs ctxt [] ctr;
val disc_concl = betapply (disc, lhs);
val (eqn_data_disc_opt, matchedsss') =
@@ -799,8 +773,7 @@
val sel_concls = sels ~~ ctr_args
|> map (fn (sel, ctr_arg) => HOLogic.mk_eq (betapply (sel, lhs), ctr_arg))
- handle ListPair.UnequalLengths =>
- error_at ctxt [rhs] "Partially applied constructor in right-hand side";
+ handle ListPair.UnequalLengths => partially_applied_ctr_in_rhs ctxt [rhs];
val eqns_data_sel =
map (dissect_coeqn_sel ctxt fun_names basic_ctr_specss eqn_pos
@@ -822,7 +795,7 @@
val cond_ctrs = fold_rev_corec_code_rhs ctxt (fn cs => fn ctr => fn _ =>
if member (op = o apsnd #ctr) basic_ctr_specs ctr then cons (ctr, cs)
- else error ("Not a constructor: " ^ quote (Syntax.string_of_term ctxt ctr))) [] rhs' []
+ else not_constructor_in_rhs ctxt [] ctr) [] rhs' []
|> AList.group (op =);
val ctr_premss = (case cond_ctrs of [_] => [[]] | _ => map (s_dnf o snd) cond_ctrs);
@@ -834,7 +807,7 @@
|> curry HOLogic.mk_eq lhs);
val bads = maps (filter (Term.exists_subterm (is_free_in fun_names))) ctr_premss;
- val _ = null bads orelse unexpected_corec_call ctxt [eqn0] rhs;
+ val _ = null bads orelse unexpected_corec_call_in ctxt [eqn0] rhs;
val sequentials = replicate (length fun_names) false;
in
@@ -847,10 +820,10 @@
(basic_ctr_specss : basic_corec_ctr_spec list list) (eqn_pos, eqn0) of_spec_opt matchedsss =
let
val eqn = drop_all eqn0
- handle TERM _ => error_at ctxt [eqn0] "Ill-formed formula";
+ handle TERM _ => ill_formed_formula ctxt eqn0;
val (prems, concl) = Logic.strip_horn eqn
|> map_prod (map HOLogic.dest_Trueprop) HOLogic.dest_Trueprop
- handle TERM _ => error_at ctxt [eqn] "Ill-formed equation";
+ handle TERM _ => ill_formed_equation ctxt eqn;
val head = concl
|> perhaps (try HOLogic.dest_not) |> perhaps (try (fst o HOLogic.dest_eq))
@@ -860,7 +833,7 @@
fun check_num_args () =
is_none rhs_opt orelse not (can dest_funT (fastype_of (the rhs_opt))) orelse
- error_at ctxt [eqn] "Expected more arguments to function on left-hand side";
+ missing_args_to_fun_on_lhs ctxt [eqn];
val discs = maps (map #disc) basic_ctr_specss;
val sels = maps (maps #sels) basic_ctr_specss;
@@ -940,8 +913,8 @@
let
val U2 =
(case try dest_sumT U of
- SOME (U1, U2) => if U1 = T then U2 else invalid_map ctxt t0
- | NONE => invalid_map ctxt t0);
+ SOME (U1, U2) => if U1 = T then U2 else invalid_map ctxt [] t0
+ | NONE => invalid_map ctxt [] t0);
fun rewrite bound_Ts (Abs (s, T', t')) = Abs (s, T', rewrite (T' :: bound_Ts) t')
| rewrite bound_Ts (t as _ $ _) =
@@ -962,7 +935,7 @@
end;
fun massage_noncall U T t =
- build_map ctxt [] (uncurry Inl_const o dest_sumT o snd) (T, U) $ t;
+ build_map ctxt [] [] (uncurry Inl_const o dest_sumT o snd) (T, U) $ t;
val bound_Ts = List.rev (map fastype_of fun_args);
in
@@ -1009,7 +982,7 @@
val bad = fold (add_extra_frees ctxt [] []) corec_args [];
val _ = null bad orelse
(if exists has_call corec_args then nonprimitive_corec ctxt []
- else extra_variable ctxt [] (hd bad));
+ else extra_variable_in_rhs ctxt [] (hd bad));
val excludess' =
disc_eqnss
@@ -1082,16 +1055,13 @@
fun is_trivial_implies thm =
uncurry (member (op aconv)) (Logic.strip_horn (Thm.prop_of thm));
-fun primcorec_ursive auto opts fixes specs of_specs_opt lthy =
+fun primcorec_ursive int auto opts fixes specs of_specs_opt lthy =
let
- val thy = Proof_Context.theory_of lthy;
-
val (bs, mxs) = map_split (apfst fst) fixes;
val (arg_Ts, res_Ts) = map (strip_type o snd o fst #>> mk_tupleT_balanced) fixes |> split_list;
- val _ = (case filter_out (fn (_, T) => Sign.of_sort thy (T, @{sort type})) (bs ~~ arg_Ts) of
- [] => ()
- | (b, _) :: _ => error ("Type of " ^ Binding.print b ^ " contains top sort"));
+ val _ = check_duplicate_const_names bs;
+ val _ = List.app (uncurry (check_top_sort lthy)) (bs ~~ arg_Ts);
val actual_nn = length bs;
@@ -1114,7 +1084,7 @@
val missing = fun_names
|> filter (map (fn Disc x => #fun_name x | Sel x => #fun_name x) eqns_data
|> not oo member (op =));
- val _ = null missing orelse error ("Missing equations for " ^ commas missing);
+ val _ = null missing orelse missing_equations_for_const (hd missing);
val callssss =
map_filter (try (fn Sel x => x)) eqns_data
@@ -1581,32 +1551,33 @@
end)
end;
- fun after_qed thmss' = fold_map Local_Theory.define defs #-> prove thmss';
+ fun after_qed thmss' =
+ fold_map Local_Theory.define defs
+ #> tap (uncurry (print_def_consts int))
+ #-> prove thmss';
in
(goalss, after_qed, lthy)
end;
-fun primcorec_ursive_cmd auto opts (raw_fixes, raw_specs_of) lthy =
+fun primcorec_ursive_cmd int auto opts (raw_fixes, raw_specs_of) lthy =
let
- val dups = duplicates (op =) (map (Binding.name_of o #1) raw_fixes);
- val _ = null dups orelse error ("Duplicate function name " ^ quote (hd dups));
-
val (raw_specs, of_specs_opt) =
split_list raw_specs_of ||> map (Option.map (Syntax.read_term lthy));
val (fixes, specs) =
fst (Specification.read_multi_specs raw_fixes (map (fn spec => (spec, [], [])) raw_specs) lthy);
in
- primcorec_ursive auto opts fixes specs of_specs_opt lthy
+ primcorec_ursive int auto opts fixes specs of_specs_opt lthy
end;
-val primcorecursive_cmd = (fn (goalss, after_qed, lthy) =>
- lthy
- |> Proof.theorem NONE after_qed goalss
- |> Proof.refine_singleton (Method.primitive_text (K I))) ooo primcorec_ursive_cmd false;
+fun primcorecursive_cmd int = (fn (goalss, after_qed, lthy) =>
+ lthy
+ |> Proof.theorem NONE after_qed goalss
+ |> Proof.refine_singleton (Method.primitive_text (K I))) ooo
+ primcorec_ursive_cmd int false;
-val primcorec_cmd = (fn (goalss, after_qed, lthy) =>
+fun primcorec_cmd int = (fn (goalss, after_qed, lthy) =>
lthy |> after_qed (map (fn [] => [] | _ => use_primcorecursive ()) goalss)) ooo
- primcorec_ursive_cmd true;
+ primcorec_ursive_cmd int true;
val corec_option_parser = Parse.group (K "option")
(Plugin_Name.parse_filter >> Plugins_Option
@@ -1621,13 +1592,13 @@
"define primitive corecursive functions"
((Scan.optional (@{keyword "("} |--
Parse.!!! (Parse.list1 corec_option_parser) --| @{keyword ")"}) []) --
- (Parse.vars -- where_alt_props_of_parser) >> uncurry primcorecursive_cmd);
+ (Parse.vars -- where_alt_props_of_parser) >> uncurry (primcorecursive_cmd true));
val _ = Outer_Syntax.local_theory @{command_keyword primcorec}
"define primitive corecursive functions"
((Scan.optional (@{keyword "("} |-- Parse.!!! (Parse.list1 corec_option_parser)
--| @{keyword ")"}) []) --
- (Parse.vars -- where_alt_props_of_parser) >> uncurry primcorec_cmd);
+ (Parse.vars -- where_alt_props_of_parser) >> uncurry (primcorec_cmd true));
val _ = Theory.setup (gfp_rec_sugar_interpretation transfer_plugin
gfp_rec_sugar_transfer_interpretation);
--- a/src/HOL/Tools/BNF/bnf_gfp_rec_sugar_tactics.ML Wed Jan 04 14:26:08 2017 +0100
+++ b/src/HOL/Tools/BNF/bnf_gfp_rec_sugar_tactics.ML Wed Jan 04 14:26:19 2017 +0100
@@ -37,8 +37,6 @@
val split_connectI = @{thms allI impI conjI};
val unfold_lets = @{thms Let_def[abs_def] split_beta}
-fun clean_blast_tac ctxt = blast_tac (put_claset (claset_of @{theory_context HOL}) ctxt);
-
fun exhaust_inst_as_projs ctxt frees thm =
let
val num_frees = length frees;
--- a/src/HOL/Tools/BNF/bnf_lfp_compat.ML Wed Jan 04 14:26:08 2017 +0100
+++ b/src/HOL/Tools/BNF/bnf_lfp_compat.ML Wed Jan 04 14:26:19 2017 +0100
@@ -208,15 +208,17 @@
let
val thy = Proof_Context.theory_of lthy;
- fun not_datatype s = error (quote s ^ " is not a datatype");
+ fun not_datatype_name s =
+ error (quote s ^ " is not a datatype");
fun not_mutually_recursive ss =
error ("{" ^ commas ss ^ "} is not a complete set of mutually recursive datatypes");
fun checked_fp_sugar_of s =
(case fp_sugar_of lthy s of
SOME (fp_sugar as {fp, fp_co_induct_sugar = SOME _, ...}) =>
- if member (op =) prefs Include_GFPs orelse fp = Least_FP then fp_sugar else not_datatype s
- | _ => not_datatype s);
+ if member (op =) prefs Include_GFPs orelse fp = Least_FP then fp_sugar
+ else not_datatype_name s
+ | _ => not_datatype_name s);
val fpTs0 as Type (_, var_As) :: _ =
map (#T o checked_fp_sugar_of o fst o dest_Type)
@@ -317,7 +319,7 @@
|>> `(fn (inducts', induct', _, rec'_thmss) =>
SOME ((inducts', induct', mk_induct_attrs ctrss'), (rec'_thmss, [])))
else
- not_datatype fpT_name1;
+ not_datatype_name fpT_name1;
val rec'_names = map (fst o dest_Const) recs';
val rec'_thms = flat rec'_thmss;
@@ -538,11 +540,11 @@
fun old_of_new f (ts, _, simpss) = (ts, f simpss);
-val primrec = apfst (old_of_new flat) ooo BNF_LFP_Rec_Sugar.primrec [];
-val primrec_global = apfst (old_of_new flat) ooo BNF_LFP_Rec_Sugar.primrec_global [];
-val primrec_overloaded = apfst (old_of_new flat) oooo BNF_LFP_Rec_Sugar.primrec_overloaded [];
+val primrec = apfst (old_of_new flat) ooo BNF_LFP_Rec_Sugar.primrec false [];
+val primrec_global = apfst (old_of_new flat) ooo BNF_LFP_Rec_Sugar.primrec_global false [];
+val primrec_overloaded = apfst (old_of_new flat) oooo BNF_LFP_Rec_Sugar.primrec_overloaded false [];
val primrec_simple = apfst (apfst fst o apsnd (old_of_new (flat o snd))) ooo
- BNF_LFP_Rec_Sugar.primrec_simple;
+ BNF_LFP_Rec_Sugar.primrec_simple false;
val _ =
Outer_Syntax.local_theory @{command_keyword datatype_compat}
--- a/src/HOL/Tools/BNF/bnf_lfp_countable.ML Wed Jan 04 14:26:08 2017 +0100
+++ b/src/HOL/Tools/BNF/bnf_lfp_countable.ML Wed Jan 04 14:26:19 2017 +0100
@@ -27,7 +27,7 @@
HEADGOAL (resolve_tac ctxt [nchotomy RS @{thm all_reg[rotated]}] THEN'
REPEAT_ALL_NEW (resolve_tac ctxt [allI, impI] ORELSE' eresolve_tac ctxt [exE, disjE]));
-fun meta_spec_mp_tac ctxt 0 = K all_tac
+fun meta_spec_mp_tac _ 0 = K all_tac
| meta_spec_mp_tac ctxt depth =
dtac ctxt meta_spec THEN' meta_spec_mp_tac ctxt (depth - 1) THEN'
dtac ctxt meta_mp THEN' assume_tac ctxt;
@@ -106,7 +106,7 @@
NONE
else if exists_subtype_in fpTs T then
let val U = mk_U T in
- SOME (mk_to_nat_checked U $ (build_map ctxt [] (snd_const o fst) (T, U) $ Bound j))
+ SOME (mk_to_nat_checked U $ (build_map ctxt [] [] (snd_const o fst) (T, U) $ Bound j))
end
else
SOME (mk_to_nat_checked T $ Bound j);
@@ -127,13 +127,14 @@
fun derive_encode_injectives_thms _ [] = []
| derive_encode_injectives_thms ctxt fpT_names0 =
let
- fun not_datatype s = error (quote s ^ " is not a datatype");
+ fun not_datatype_name s =
+ error (quote s ^ " is not a datatype");
fun not_mutually_recursive ss = error (commas ss ^ " are not mutually recursive datatypes");
fun lfp_sugar_of s =
(case fp_sugar_of ctxt s of
SOME (fp_sugar as {fp = Least_FP, fp_co_induct_sugar = SOME _, ...}) => fp_sugar
- | _ => not_datatype s);
+ | _ => not_datatype_name s);
val fpTs0 as Type (_, var_As) :: _ =
map (#T o lfp_sugar_of o fst o dest_Type) (#Ts (#fp_res (lfp_sugar_of (hd fpT_names0))));
--- a/src/HOL/Tools/BNF/bnf_lfp_rec_sugar.ML Wed Jan 04 14:26:08 2017 +0100
+++ b/src/HOL/Tools/BNF/bnf_lfp_rec_sugar.ML Wed Jan 04 14:26:19 2017 +0100
@@ -62,21 +62,20 @@
val lfp_rec_sugar_interpretation: string ->
(BNF_FP_Rec_Sugar_Util.fp_rec_sugar -> local_theory -> local_theory) -> theory -> theory
- val primrec: rec_option list -> (binding * typ option * mixfix) list ->
+ val primrec: bool -> rec_option list -> (binding * typ option * mixfix) list ->
Specification.multi_specs -> local_theory ->
(term list * thm list * thm list list) * local_theory
- val primrec_cmd: rec_option list -> (binding * string option * mixfix) list ->
+ val primrec_cmd: bool -> rec_option list -> (binding * string option * mixfix) list ->
Specification.multi_specs_cmd -> local_theory ->
(term list * thm list * thm list list) * local_theory
- val primrec_global: rec_option list -> (binding * typ option * mixfix) list ->
+ val primrec_global: bool -> rec_option list -> (binding * typ option * mixfix) list ->
Specification.multi_specs -> theory -> (term list * thm list * thm list list) * theory
- val primrec_overloaded: rec_option list -> (string * (string * typ) * bool) list ->
+ val primrec_overloaded: bool -> rec_option list -> (string * (string * typ) * bool) list ->
(binding * typ option * mixfix) list ->
Specification.multi_specs -> theory -> (term list * thm list * thm list list) * theory
- val primrec_simple: ((binding * typ) * mixfix) list -> term list ->
- local_theory ->
- ((string list * (binding -> binding) list) *
- (term list * thm list * (int list list * thm list list))) * local_theory
+ val primrec_simple: bool -> ((binding * typ) * mixfix) list -> term list -> local_theory ->
+ ((string list * (binding -> binding) list)
+ * (term list * thm list * (int list list * thm list list))) * local_theory
end;
structure BNF_LFP_Rec_Sugar : BNF_LFP_REC_SUGAR =
@@ -96,11 +95,6 @@
exception OLD_PRIMREC of unit;
-fun malformed_eqn_lhs_rhs ctxt eqn =
- error_at ctxt [eqn] "Malformed equation (expected \"lhs = rhs\")";
-fun malformed_eqn_head ctxt eqn =
- error_at ctxt [eqn] "Malformed function equation (expected function name on left-hand side)";
-
datatype rec_option =
Plugins_Option of Proof.context -> Plugin_Name.filter |
Nonexhaustive_Option |
@@ -179,6 +173,8 @@
in
([], [0], [basic_lfp_sugar], [], [], [], TrueI (*dummy*), [], false, ctxt)
end
+ | default_basic_lfp_sugars_of _ [T] _ _ ctxt =
+ error ("Cannot recurse through type " ^ quote (Syntax.string_of_typ ctxt T))
| default_basic_lfp_sugars_of _ _ _ _ _ = error "Unsupported mutual recursion at this stage";
fun basic_lfp_sugars_of bs arg_Ts callers callssss lthy =
@@ -288,29 +284,25 @@
fun dissect_eqn ctxt fun_names eqn0 =
let
val eqn = drop_all eqn0 |> HOLogic.dest_Trueprop
- handle TERM _ => malformed_eqn_lhs_rhs ctxt eqn0;
+ handle TERM _ => ill_formed_equation_lhs_rhs ctxt [eqn0];
val (lhs, rhs) = HOLogic.dest_eq eqn
- handle TERM _ => malformed_eqn_lhs_rhs ctxt eqn;
+ handle TERM _ => ill_formed_equation_lhs_rhs ctxt [eqn];
val (fun_name, args) = strip_comb lhs
- |>> (fn x => if is_Free x then fst (dest_Free x) else malformed_eqn_head ctxt eqn);
+ |>> (fn x => if is_Free x then fst (dest_Free x) else ill_formed_equation_head ctxt [eqn]);
val (left_args, rest) = take_prefix is_Free args;
val (nonfrees, right_args) = take_suffix is_Free rest;
val num_nonfrees = length nonfrees;
val _ = num_nonfrees = 1 orelse
- error_at ctxt [eqn] (if num_nonfrees = 0 then "Constructor pattern missing in left-hand side"
- else "More than one non-variable argument in left-hand side");
- val _ = member (op =) fun_names fun_name orelse raise malformed_eqn_head ctxt eqn;
+ (if num_nonfrees = 0 then missing_pattern ctxt [eqn]
+ else more_than_one_nonvar_in_lhs ctxt [eqn]);
+ val _ = member (op =) fun_names fun_name orelse raise ill_formed_equation_head ctxt [eqn];
val (ctr, ctr_args) = strip_comb (the_single nonfrees);
val _ = try (num_binder_types o fastype_of) ctr = SOME (length ctr_args) orelse
- error_at ctxt [eqn] "Partially applied constructor in pattern";
+ partially_applied_ctr_in_pattern ctxt [eqn];
- val dups = duplicates (op =) (left_args @ ctr_args @ right_args)
- val _ = null dups orelse
- error_at ctxt [eqn] ("Duplicate variable \"" ^ Syntax.string_of_term ctxt (hd dups) ^
- "\" in left-hand side");
- val _ = forall is_Free ctr_args orelse
- error_at ctxt [eqn] "Nonprimitive pattern in left-hand side";
+ val _ = check_duplicate_variables_in_lhs ctxt [eqn] (left_args @ ctr_args @ right_args)
+ val _ = forall is_Free ctr_args orelse nonprimitive_pattern_in_lhs ctxt [eqn];
val _ =
let
val bads =
@@ -322,8 +314,7 @@
I
| _ => I) rhs [];
in
- null bads orelse
- error_at ctxt [eqn] ("Extra variable " ^ quote (Syntax.string_of_term ctxt (hd bads)))
+ null bads orelse extra_variable_in_rhs ctxt [eqn] (hd bads)
end;
in
{fun_name = fun_name,
@@ -360,9 +351,7 @@
(case try (get_ctr_pos o fst o dest_Free) g of
SOME ~1 => subst_comb t
| SOME ctr_pos =>
- (length g_args >= ctr_pos orelse
- error ("Too few arguments in recursive call " ^
- quote (Syntax.string_of_term ctxt t));
+ (length g_args >= ctr_pos orelse too_few_args_in_rec_call ctxt [] t;
(case AList.lookup (op =) mutual_calls y of
SOME y' => list_comb (y', map (subst bound_Ts) g_args)
| NONE => subst_comb t))
@@ -373,11 +362,8 @@
| subst _ t = t
fun subst' t =
- if has_call t then
- error ("Recursive call not directly applied to constructor argument in " ^
- quote (Syntax.string_of_term ctxt t))
- else
- try_nested_rec [] (head_of t) t |> the_default t;
+ if has_call t then rec_call_not_apply_to_ctr_arg ctxt [] t
+ else try_nested_rec [] (head_of t) t |> the_default t;
in
subst' o subst []
end;
@@ -436,20 +422,20 @@
let
val zs = replicate (length xs) z;
val (b, c) = finds (fn ((x, _), y) => #ctr x = #ctr y) (xs ~~ zs) ys;
- val _ = null c orelse error_at ctxt (map #rhs_term c) "Excess equation(s)";
+ val _ = null c orelse excess_equations ctxt (map #rhs_term c);
in b end) (map #ctr_specs (take n_funs rec_specs) ~~ funs_data ~~ nonexhaustives);
val (_ : unit list) = ctr_spec_eqn_data_list' |> map (fn (({ctr, ...}, nonexhaustive), x) =>
if length x > 1 then
- error_at ctxt (map #user_eqn x) "Multiple equations for constructor"
+ multiple_equations_for_ctr ctxt (map #user_eqn x)
else if length x = 1 orelse nonexhaustive orelse not (Context_Position.is_visible ctxt) then
()
else
- warning ("No equation for constructor " ^ Syntax.string_of_term ctxt ctr));
+ no_equation_for_ctr_warning ctxt [] ctr);
val ctr_spec_eqn_data_list =
map (apfst fst) ctr_spec_eqn_data_list' @
- (drop n_funs rec_specs |> maps #ctr_specs |> map (rpair []));
+ (drop n_funs rec_specs |> maps #ctr_specs |> map (rpair []));
val recs = take n_funs rec_specs |> map #recx;
val rec_args = ctr_spec_eqn_data_list
@@ -457,7 +443,7 @@
|> map (uncurry (build_rec_arg ctxt funs_data has_call) o apsnd (try the_single));
val ctr_poss = map (fn x =>
if length (distinct (op = o apply2 (length o #left_args)) x) <> 1 then
- error ("Inconstant constructor pattern position for function " ^ quote (#fun_name (hd x)))
+ inconstant_pattern_pos_for_fun ctxt [] (#fun_name (hd x))
else
hd x |> #left_args |> length) funs_data;
in
@@ -518,7 +504,7 @@
|> partition_eq (op = o apply2 #fun_name)
|> finds (fn (x, y) => x = #fun_name (hd y)) fun_names |> fst
|> map (fn (x, y) => the_single y
- handle List.Empty => error ("Missing equations for function " ^ quote x));
+ handle List.Empty => missing_equations_for_fun x);
val frees = map (fst #>> Binding.name_of #> Free) fixes;
val has_call = exists_subterm (member (op =) frees);
@@ -533,9 +519,7 @@
| is_only_old_datatype _ = false;
val _ = if exists is_only_old_datatype arg_Ts then raise OLD_PRIMREC () else ();
- val _ = (case filter_out (fn (_, T) => Sign.of_sort thy (T, @{sort type})) (bs ~~ res_Ts) of
- [] => ()
- | (b, _) :: _ => error ("Type of " ^ Binding.print b ^ " contains top sort"));
+ val _ = List.app (uncurry (check_top_sort lthy0)) (bs ~~ res_Ts);
val ((n2m, rec_specs, _, common_induct, inducts, induct_attrs, Ts), lthy) =
rec_specs_of bs arg_Ts res_Ts frees callssss lthy0;
@@ -543,10 +527,9 @@
val actual_nn = length funs_data;
val ctrs = maps (map #ctr o #ctr_specs) rec_specs;
- val _ =
- map (fn {ctr, user_eqn, ...} => member (op =) ctrs ctr orelse
- error_at lthy0 [user_eqn] ("Argument " ^ quote (Syntax.string_of_term lthy ctr) ^
- " is not a constructor in left-hand side")) eqns_data;
+ val _ = List.app (fn {ctr, user_eqn, ...} =>
+ ignore (member (op =) ctrs ctr orelse not_constructor_in_pattern lthy0 [user_eqn] ctr))
+ eqns_data;
val defs = build_defs lthy nonexhaustives bs mxs funs_data rec_specs has_call;
@@ -605,8 +588,10 @@
lthy |> Local_Theory.notes (notes @ common_notes) |> snd)
end;
-fun primrec_simple0 plugins nonexhaustive transfer fixes ts lthy =
+fun primrec_simple0 int plugins nonexhaustive transfer fixes ts lthy =
let
+ val _ = check_duplicate_const_names (map (fst o fst) fixes);
+
val actual_nn = length fixes;
val nonexhaustives = replicate actual_nn nonexhaustive;
@@ -617,24 +602,22 @@
in
lthy'
|> fold_map Local_Theory.define defs
+ |> tap (uncurry (print_def_consts int))
|-> (fn defs => fn lthy =>
let val ((jss, simpss), lthy) = prove lthy defs in
(((names, qualifys), (map fst defs, map (snd o snd) defs, (jss, simpss))), lthy)
end)
end;
-fun primrec_simple fixes ts lthy =
- primrec_simple0 Plugin_Name.default_filter false false fixes ts lthy
+fun primrec_simple int fixes ts lthy =
+ primrec_simple0 int Plugin_Name.default_filter false false fixes ts lthy
handle OLD_PRIMREC () =>
- Old_Primrec.primrec_simple fixes ts lthy
+ Old_Primrec.primrec_simple int fixes ts lthy
|>> apsnd (fn (ts, thms) => (ts, [], ([], [thms]))) o apfst single
|>> apfst (map_split (rpair I));
-fun gen_primrec old_primrec prep_spec opts raw_fixes raw_specs lthy =
+fun gen_primrec old_primrec prep_spec int opts raw_fixes raw_specs lthy =
let
- val dups = duplicates (op =) (map (fn (x, _, _) => Binding.name_of x) raw_fixes);
- val _ = null dups orelse error ("Duplicate function name " ^ quote (hd dups));
-
val plugins = get_first (fn Plugins_Option f => SOME (f lthy) | _ => NONE) (rev opts)
|> the_default Plugin_Name.default_filter;
val nonexhaustive = exists (can (fn Nonexhaustive_Option => ())) opts;
@@ -658,27 +641,27 @@
end);
in
lthy
- |> primrec_simple0 plugins nonexhaustive transfer fixes (map snd specs)
+ |> primrec_simple0 int plugins nonexhaustive transfer fixes (map snd specs)
|-> (fn ((names, qualifys), (ts, defs, (jss, simpss))) =>
Spec_Rules.add Spec_Rules.Equational (ts, flat simpss)
#> Local_Theory.notes (mk_notes jss names qualifys simpss)
#>> (fn notes => (ts, defs, map_filter (fn ("", _) => NONE | (_, thms) => SOME thms) notes)))
end
handle OLD_PRIMREC () =>
- old_primrec raw_fixes raw_specs lthy
+ old_primrec int raw_fixes raw_specs lthy
|>> (fn (ts, thms) => (ts, [], [thms]));
val primrec = gen_primrec Old_Primrec.primrec Specification.check_multi_specs;
val primrec_cmd = gen_primrec Old_Primrec.primrec_cmd Specification.read_multi_specs;
-fun primrec_global opts fixes specs =
+fun primrec_global int opts fixes specs =
Named_Target.theory_init
- #> primrec opts fixes specs
+ #> primrec int opts fixes specs
##> Local_Theory.exit_global;
-fun primrec_overloaded opts ops fixes specs =
+fun primrec_overloaded int opts ops fixes specs =
Overloading.overloading ops
- #> primrec opts fixes specs
+ #> primrec int opts fixes specs
##> Local_Theory.exit_global;
val rec_option_parser = Parse.group (K "option")
@@ -690,6 +673,6 @@
"define primitive recursive functions"
((Scan.optional (@{keyword "("} |-- Parse.!!! (Parse.list1 rec_option_parser)
--| @{keyword ")"}) []) -- Parse_Spec.specification
- >> (fn (opts, (fixes, specs)) => snd o primrec_cmd opts fixes specs));
+ >> (fn (opts, (fixes, specs)) => snd o primrec_cmd true opts fixes specs));
end;
--- a/src/HOL/Tools/BNF/bnf_lfp_rec_sugar_more.ML Wed Jan 04 14:26:08 2017 +0100
+++ b/src/HOL/Tools/BNF/bnf_lfp_rec_sugar_more.ML Wed Jan 04 14:26:19 2017 +0100
@@ -74,19 +74,12 @@
exception NO_MAP of term;
-fun ill_formed_rec_call ctxt t =
- error ("Ill-formed recursive call: " ^ quote (Syntax.string_of_term ctxt t));
-fun invalid_map ctxt t =
- error ("Invalid map function in " ^ quote (Syntax.string_of_term ctxt t));
-fun unexpected_rec_call ctxt eqns t =
- error_at ctxt eqns ("Unexpected recursive call in " ^ quote (Syntax.string_of_term ctxt t));
-
fun massage_nested_rec_call ctxt has_call massage_fun massage_nonfun bound_Ts y y' t0 =
let
- fun check_no_call t = if has_call t then unexpected_rec_call ctxt [t0] t else ();
+ fun check_no_call t = if has_call t then unexpected_rec_call_in ctxt [t0] t else ();
val typof = curry fastype_of1 bound_Ts;
- val massage_no_call = build_map ctxt [] massage_nonfun;
+ val massage_no_call = build_map ctxt [] [] massage_nonfun;
val yT = typof y;
val yU = typof y';
@@ -133,11 +126,11 @@
if has_call t then
if t2 = y then
massage_map yU yT (elim_y t1) $ y'
- handle NO_MAP t' => invalid_map ctxt t'
+ handle NO_MAP t' => invalid_map ctxt [t0] t'
else
let val (g, xs) = Term.strip_comb t2 in
if g = y then
- if exists has_call xs then unexpected_rec_call ctxt [t0] t2
+ if exists has_call xs then unexpected_rec_call_in ctxt [t0] t2
else Term.list_comb (massage_outer_call (mk_compN (length xs) bound_Ts (t1, y)), xs)
else
ill_formed_rec_call ctxt t
@@ -153,8 +146,8 @@
let
val _ =
(case try HOLogic.dest_prodT U of
- SOME (U1, _) => U1 = T orelse invalid_map ctxt t
- | NONE => invalid_map ctxt t);
+ SOME (U1, _) => U1 = T orelse invalid_map ctxt [] t
+ | NONE => invalid_map ctxt [] t);
fun subst d (t as Bound d') = t |> d = SOME d' ? curry (op $) (fst_const U)
| subst d (Abs (v, T, b)) =
@@ -171,8 +164,7 @@
d = try (fn Bound n => n) (nth vs ctr_pos) then
Term.list_comb (snd_const U $ nth vs ctr_pos, map (subst d) (nth_drop ctr_pos vs))
else
- error ("Recursive call not directly applied to constructor argument in " ^
- quote (Syntax.string_of_term ctxt t))
+ rec_call_not_apply_to_ctr_arg ctxt [] t
else
Term.list_comb (u, map (subst (if d = SOME ~1 then NONE else d)) vs)
end;
--- a/src/HOL/Tools/BNF/bnf_tactics.ML Wed Jan 04 14:26:08 2017 +0100
+++ b/src/HOL/Tools/BNF/bnf_tactics.ML Wed Jan 04 14:26:19 2017 +0100
@@ -11,6 +11,7 @@
include CTR_SUGAR_GENERAL_TACTICS
val fo_rtac: Proof.context -> thm -> int -> tactic
+ val clean_blast_tac: Proof.context -> int -> tactic
val subst_tac: Proof.context -> int list option -> thm list -> int -> tactic
val subst_asm_tac: Proof.context -> int list option -> thm list -> int -> tactic
@@ -43,6 +44,8 @@
end
handle Pattern.MATCH => no_tac) ctxt;
+fun clean_blast_tac ctxt = blast_tac (put_claset (claset_of @{theory_context HOL}) ctxt);
+
(*unlike "unfold_thms_tac", it succeed when the RHS contains schematic variables not in the LHS*)
fun subst_tac ctxt = EqSubst.eqsubst_tac ctxt o the_default [0];
fun subst_asm_tac ctxt = EqSubst.eqsubst_asm_tac ctxt o the_default [0];
--- a/src/HOL/Tools/Old_Datatype/old_primrec.ML Wed Jan 04 14:26:08 2017 +0100
+++ b/src/HOL/Tools/Old_Datatype/old_primrec.ML Wed Jan 04 14:26:19 2017 +0100
@@ -8,16 +8,16 @@
signature OLD_PRIMREC =
sig
- val primrec: (binding * typ option * mixfix) list ->
+ val primrec: bool -> (binding * typ option * mixfix) list ->
Specification.multi_specs -> local_theory -> (term list * thm list) * local_theory
- val primrec_cmd: (binding * string option * mixfix) list ->
+ val primrec_cmd: bool -> (binding * string option * mixfix) list ->
Specification.multi_specs_cmd -> local_theory -> (term list * thm list) * local_theory
- val primrec_global: (binding * typ option * mixfix) list ->
+ val primrec_global: bool -> (binding * typ option * mixfix) list ->
Specification.multi_specs -> theory -> (term list * thm list) * theory
- val primrec_overloaded: (string * (string * typ) * bool) list ->
+ val primrec_overloaded: bool -> (string * (string * typ) * bool) list ->
(binding * typ option * mixfix) list ->
Specification.multi_specs -> theory -> (term list * thm list) * theory
- val primrec_simple: ((binding * typ) * mixfix) list -> term list ->
+ val primrec_simple: bool -> ((binding * typ) * mixfix) list -> term list ->
local_theory -> (string * (term list * thm list)) * local_theory
end;
@@ -188,12 +188,12 @@
in
(dummy_fns @ fs, defs)
end
- | SOME (fname, ls, fs') => (fs' @ fs, (fname, ls, rec_name, tname) :: defs));
+ | SOME (fname, ls, fs') => (fs' @ fs, (fname, ls, rec_name) :: defs));
(* make definition *)
-fun make_def ctxt fixes fs (fname, ls, rec_name, tname) =
+fun make_def ctxt fixes fs (fname, ls, rec_name) =
let
val SOME (var, varT) = get_first (fn ((b, T), mx) =>
if Binding.name_of b = fname then SOME ((b, mx), T) else NONE) fixes;
@@ -260,18 +260,19 @@
(* primrec definition *)
-fun primrec_simple fixes ts lthy =
+fun primrec_simple int fixes ts lthy =
let
val ((prefix, (_, defs)), prove) = distill lthy fixes ts;
in
lthy
|> fold_map Local_Theory.define defs
+ |> tap (uncurry (BNF_FP_Rec_Sugar_Util.print_def_consts int))
|-> (fn defs => `(fn lthy => (prefix, (map fst defs, prove lthy defs))))
end;
local
-fun gen_primrec prep_spec raw_fixes raw_spec lthy =
+fun gen_primrec prep_spec int raw_fixes raw_spec lthy =
let
val (fixes, spec) = fst (prep_spec raw_fixes raw_spec lthy);
fun attr_bindings prefix = map (fn ((b, attrs), _) =>
@@ -280,7 +281,7 @@
(Binding.qualify true prefix (Binding.name "simps"), @{attributes [simp, nitpick_simp]});
in
lthy
- |> primrec_simple fixes (map snd spec)
+ |> primrec_simple int fixes (map snd spec)
|-> (fn (prefix, (ts, simps)) =>
Spec_Rules.add Spec_Rules.Equational (ts, simps)
#> fold_map Local_Theory.note (attr_bindings prefix ~~ map single simps)
@@ -295,17 +296,17 @@
end;
-fun primrec_global fixes specs thy =
+fun primrec_global int fixes specs thy =
let
val lthy = Named_Target.theory_init thy;
- val ((ts, simps), lthy') = primrec fixes specs lthy;
+ val ((ts, simps), lthy') = primrec int fixes specs lthy;
val simps' = Proof_Context.export lthy' lthy simps;
in ((ts, simps'), Local_Theory.exit_global lthy') end;
-fun primrec_overloaded ops fixes specs thy =
+fun primrec_overloaded int ops fixes specs thy =
let
val lthy = Overloading.overloading ops thy;
- val ((ts, simps), lthy') = primrec fixes specs lthy;
+ val ((ts, simps), lthy') = primrec int fixes specs lthy;
val simps' = Proof_Context.export lthy' lthy simps;
in ((ts, simps'), Local_Theory.exit_global lthy') end;
--- a/src/HOL/Tools/Predicate_Compile/predicate_compile_fun.ML Wed Jan 04 14:26:08 2017 +0100
+++ b/src/HOL/Tools/Predicate_Compile/predicate_compile_fun.ML Wed Jan 04 14:26:19 2017 +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/Predicate_Compile/predicate_compile_proof.ML Wed Jan 04 14:26:08 2017 +0100
+++ b/src/HOL/Tools/Predicate_Compile/predicate_compile_proof.ML Wed Jan 04 14:26:19 2017 +0100
@@ -152,6 +152,7 @@
THEN TRY (
let
val prems' = maps dest_conjunct_prem (take nargs prems)
+ |> filter is_equationlike
in
rewrite_goal_tac ctxt'
(map (fn th => th RS @{thm sym} RS @{thm eq_reflection}) prems') 1
@@ -314,7 +315,7 @@
THEN (DETERM (TRY (etac @{thm Pair_inject} 1)))*)
THEN (REPEAT_DETERM_N (num_of_constrs - 1)
(eresolve_tac ctxt @{thms botE} 1 ORELSE eresolve_tac ctxt @{thms botE} 2)))
- THEN assert_tac ctxt @{here} (fn st => Thm.nprems_of st <= 2)
+ THEN assert_tac ctxt \<^here> (fn st => Thm.nprems_of st <= 2)
THEN (EVERY (map split_term_tac ts))
end
else all_tac
--- a/src/HOL/Tools/Qelim/cooper.ML Wed Jan 04 14:26:08 2017 +0100
+++ b/src/HOL/Tools/Qelim/cooper.ML Wed Jan 04 14:26:19 2017 +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/Tools/Sledgehammer/sledgehammer_util.ML Wed Jan 04 14:26:08 2017 +0100
+++ b/src/HOL/Tools/Sledgehammer/sledgehammer_util.ML Wed Jan 04 14:26:19 2017 +0100
@@ -84,8 +84,10 @@
be missing over there; or maybe the two code portions are not doing the same? *)
fun fold_body_thm max_thms outer_name (map_plain_name, map_inclass_name) body =
let
- fun app_thm map_name (_, (name, _, body)) (accum as (num_thms, names)) =
+ fun app_thm map_name (_, thm_node) (accum as (num_thms, names)) =
let
+ val name = Proofterm.thm_node_name thm_node
+ val body = Proofterm.thm_node_body thm_node
val (anonymous, enter_class) =
(* The "name = outer_name" case caters for the uncommon case where the proved theorem
occurs in its own proof (e.g., "Transitive_Closure.trancl_into_trancl"). *)
--- a/src/HOL/Topological_Spaces.thy Wed Jan 04 14:26:08 2017 +0100
+++ b/src/HOL/Topological_Spaces.thy Wed Jan 04 14:26:19 2017 +0100
@@ -1547,6 +1547,18 @@
lemma tendsto_compose_filtermap: "((g \<circ> f) \<longlongrightarrow> T) F \<longleftrightarrow> (g \<longlongrightarrow> T) (filtermap f F)"
by (simp add: filterlim_def filtermap_filtermap comp_def)
+lemma tendsto_compose_at:
+ assumes f: "(f \<longlongrightarrow> y) F" and g: "(g \<longlongrightarrow> z) (at y)" and fg: "eventually (\<lambda>w. f w = y \<longrightarrow> g y = z) F"
+ shows "((g \<circ> f) \<longlongrightarrow> z) F"
+proof -
+ have "(\<forall>\<^sub>F a in F. f a \<noteq> y) \<or> g y = z"
+ using fg by force
+ moreover have "(g \<longlongrightarrow> z) (filtermap f F) \<or> \<not> (\<forall>\<^sub>F a in F. f a \<noteq> y)"
+ by (metis (no_types) filterlim_atI filterlim_def tendsto_mono f g)
+ ultimately show ?thesis
+ by (metis (no_types) f filterlim_compose filterlim_filtermap g tendsto_at_iff_tendsto_nhds tendsto_compose_filtermap)
+qed
+
subsubsection \<open>Relation of \<open>LIM\<close> and \<open>LIMSEQ\<close>\<close>
@@ -2087,12 +2099,10 @@
lemma compact_empty[simp]: "compact {}"
by (auto intro!: compactI)
-lemma compactE:
- assumes "compact s"
- and "\<forall>t\<in>C. open t"
- and "s \<subseteq> \<Union>C"
- obtains C' where "C' \<subseteq> C" and "finite C'" and "s \<subseteq> \<Union>C'"
- using assms unfolding compact_eq_heine_borel by metis
+lemma compactE: (*related to COMPACT_IMP_HEINE_BOREL in HOL Light*)
+ assumes "compact S" "S \<subseteq> \<Union>\<T>" "\<And>B. B \<in> \<T> \<Longrightarrow> open B"
+ obtains \<T>' where "\<T>' \<subseteq> \<T>" "finite \<T>'" "S \<subseteq> \<Union>\<T>'"
+ by (meson assms compact_eq_heine_borel)
lemma compactE_image:
assumes "compact s"
@@ -2197,9 +2207,7 @@
fix y
assume "y \<in> - s"
let ?C = "\<Union>x\<in>s. {u. open u \<and> x \<in> u \<and> eventually (\<lambda>y. y \<notin> u) (nhds y)}"
- note \<open>compact s\<close>
- moreover have "\<forall>u\<in>?C. open u" by simp
- moreover have "s \<subseteq> \<Union>?C"
+ have "s \<subseteq> \<Union>?C"
proof
fix x
assume "x \<in> s"
@@ -2209,8 +2217,8 @@
with \<open>x \<in> s\<close> show "x \<in> \<Union>?C"
unfolding eventually_nhds by auto
qed
- ultimately obtain D where "D \<subseteq> ?C" and "finite D" and "s \<subseteq> \<Union>D"
- by (rule compactE)
+ then obtain D where "D \<subseteq> ?C" and "finite D" and "s \<subseteq> \<Union>D"
+ by (rule compactE [OF \<open>compact s\<close>]) auto
from \<open>D \<subseteq> ?C\<close> have "\<forall>x\<in>D. eventually (\<lambda>y. y \<notin> x) (nhds y)"
by auto
with \<open>finite D\<close> have "eventually (\<lambda>y. y \<notin> \<Union>D) (nhds y)"
--- a/src/HOL/Transcendental.thy Wed Jan 04 14:26:08 2017 +0100
+++ b/src/HOL/Transcendental.thy Wed Jan 04 14:26:19 2017 +0100
@@ -2370,6 +2370,53 @@
qed (rule exp_at_top)
qed
+subsubsection\<open> A couple of simple bounds\<close>
+
+lemma exp_plus_inverse_exp:
+ fixes x::real
+ shows "2 \<le> exp x + inverse (exp x)"
+proof -
+ have "2 \<le> exp x + exp (-x)"
+ using exp_ge_add_one_self [of x] exp_ge_add_one_self [of "-x"]
+ by linarith
+ then show ?thesis
+ by (simp add: exp_minus)
+qed
+
+lemma real_le_x_sinh:
+ fixes x::real
+ assumes "0 \<le> x"
+ shows "x \<le> (exp x - inverse(exp x)) / 2"
+proof -
+ have *: "exp a - inverse(exp a) - 2*a \<le> exp b - inverse(exp b) - 2*b" if "a \<le> b" for a b::real
+ apply (rule DERIV_nonneg_imp_nondecreasing [OF that])
+ using exp_plus_inverse_exp
+ apply (intro exI allI impI conjI derivative_eq_intros | force)+
+ done
+ show ?thesis
+ using*[OF assms] by simp
+qed
+
+lemma real_le_abs_sinh:
+ fixes x::real
+ shows "abs x \<le> abs((exp x - inverse(exp x)) / 2)"
+proof (cases "0 \<le> x")
+ case True
+ show ?thesis
+ using real_le_x_sinh [OF True] True by (simp add: abs_if)
+next
+ case False
+ have "-x \<le> (exp(-x) - inverse(exp(-x))) / 2"
+ by (meson False linear neg_le_0_iff_le real_le_x_sinh)
+ also have "... \<le> \<bar>(exp x - inverse (exp x)) / 2\<bar>"
+ by (metis (no_types, hide_lams) abs_divide abs_le_iff abs_minus_cancel
+ add.inverse_inverse exp_minus minus_diff_eq order_refl)
+ finally show ?thesis
+ using False by linarith
+qed
+
+subsection\<open>The general logarithm\<close>
+
definition log :: "real \<Rightarrow> real \<Rightarrow> real"
\<comment> \<open>logarithm of @{term x} to base @{term a}\<close>
where "log a x = ln x / ln a"
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Types_To_Sets/Examples/Finite.thy Wed Jan 04 14:26:19 2017 +0100
@@ -0,0 +1,90 @@
+(* Title: HOL/Types_To_Sets/Examples/Finite.thy
+ Author: Ondřej Kunčar, TU München
+*)
+
+theory Finite
+ imports "../Types_To_Sets" Prerequisites
+begin
+
+section \<open>The Type-Based Theorem\<close>
+
+text\<open>This example file shows how we perform the relativization in presence of axiomatic
+ type classes: we internalize them.\<close>
+
+definition injective :: "('a \<Rightarrow> 'b) \<Rightarrow> bool"
+ where "injective f \<equiv> (\<forall>x y. f x = f y \<longrightarrow> x = y)"
+
+text\<open>We want to relativize the following type-based theorem using the type class finite.\<close>
+
+lemma type_based: "\<exists>f :: 'a::finite \<Rightarrow> nat. injective f"
+ unfolding injective_def
+ using finite_imp_inj_to_nat_seg[of "UNIV::'a set", unfolded inj_on_def] by auto
+
+section \<open>Definitions and Setup for The Relativization\<close>
+
+text\<open>We have to define what being injective on a set means.\<close>
+
+definition injective_on :: "'a set \<Rightarrow> ('a \<Rightarrow> 'b) \<Rightarrow> bool"
+ where "injective_on A f \<equiv> \<forall>x \<in> A. \<forall>y \<in> A. f x = f y \<longrightarrow> x = y"
+
+text\<open>The predicate injective_on is the relativization of the predicate injective.\<close>
+
+lemma injective_transfer[transfer_rule]:
+ includes lifting_syntax
+ assumes [transfer_rule]: "right_total T"
+ assumes [transfer_rule]: "bi_unique T"
+ shows "((T ===> op =) ===> op=) (injective_on (Collect(Domainp T))) injective"
+ unfolding injective_on_def[abs_def] injective_def[abs_def] by transfer_prover
+
+text\<open>Similarly, we define the relativization of the internalization
+ of the type class finite, class.finite\<close>
+
+definition finite_on :: "'a set => bool" where "finite_on A = finite A"
+
+lemma class_finite_transfer[transfer_rule]:
+ assumes [transfer_rule]: "right_total (T::'a\<Rightarrow>'b\<Rightarrow>bool)"
+ assumes [transfer_rule]: "bi_unique T"
+ shows "op= (finite_on (Collect(Domainp T))) (class.finite TYPE('b))"
+ unfolding finite_on_def class.finite_def by transfer_prover
+
+text\<open>The aforementioned development can be automated. The main part is already automated
+ by the transfer_prover.\<close>
+
+section \<open>The Relativization to The Set-Based Theorem\<close>
+
+lemmas internalized = type_based[internalize_sort "'a::finite"]
+text\<open>We internalized the type class finite and thus reduced the task to the
+ original relativization algorithm.\<close>
+thm internalized
+
+text\<open>This is the set-based variant.\<close>
+
+lemma set_based:
+ assumes "(A::'a set) \<noteq> {}"
+ shows "finite_on A \<longrightarrow> (\<exists>f :: 'a \<Rightarrow> nat. injective_on A f)"
+proof -
+ {
+ text\<open>We define the type 'b to be isomorphic to A.\<close>
+ assume T: "\<exists>(Rep :: 'b \<Rightarrow> 'a) Abs. type_definition Rep Abs A"
+ from T obtain rep :: "'b \<Rightarrow> 'a" and abs :: "'a \<Rightarrow> 'b" where t: "type_definition rep abs A"
+ by auto
+
+ text\<open>Setup for the Transfer tool.\<close>
+ define cr_b where "cr_b == \<lambda>r a. r = rep a"
+ note type_definition_Domainp[OF t cr_b_def, transfer_domain_rule]
+ note typedef_right_total[OF t cr_b_def, transfer_rule]
+ note typedef_bi_unique[OF t cr_b_def, transfer_rule]
+
+ have ?thesis
+ text\<open>Relativization by the Transfer tool.\<close>
+ using internalized[where 'a = 'b, untransferred, simplified]
+ by blast
+ } note * = this[cancel_type_definition, OF assms] text\<open>We removed the definition of 'b.\<close>
+
+ show ?thesis by (rule *)
+qed
+
+text\<open>The Final Result. We can compare the type-based and the set-based statement.\<close>
+thm type_based set_based
+
+end
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Types_To_Sets/Examples/Prerequisites.thy Wed Jan 04 14:26:19 2017 +0100
@@ -0,0 +1,25 @@
+(* Title: HOL/Types_To_Sets/Examples/Prerequisites.thy
+ Author: Ondřej Kunčar, TU München
+*)
+
+theory Prerequisites
+ imports Main
+begin
+
+context
+ fixes Rep Abs A T
+ assumes type: "type_definition Rep Abs A"
+ assumes T_def: "T \<equiv> (\<lambda>(x::'a) (y::'b). x = Rep y)"
+begin
+
+lemma type_definition_Domainp: "Domainp T = (\<lambda>x. x \<in> A)"
+proof -
+ interpret type_definition Rep Abs A by(rule type)
+ show ?thesis
+ unfolding Domainp_iff[abs_def] T_def fun_eq_iff
+ by (metis Abs_inverse Rep)
+qed
+
+end
+
+end
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Types_To_Sets/Examples/T2_Spaces.thy Wed Jan 04 14:26:19 2017 +0100
@@ -0,0 +1,162 @@
+(* Title: HOL/Types_To_Sets/Examples/T2_Spaces.thy
+ Author: Ondřej Kunčar, TU München
+*)
+
+theory T2_Spaces
+ imports Complex_Main "../Types_To_Sets" Prerequisites
+begin
+
+section \<open>The Type-Based Theorem\<close>
+
+text\<open>We relativize a theorem that contains a type class with an associated (overloaded) operation.
+ The key technique is to compile out the overloaded operation by the dictionary construction
+ using the Unoverloading rule.\<close>
+
+text\<open>This is the type-based statement that we want to relativize.\<close>
+thm compact_imp_closed
+text\<open>The type is class a T2 typological space.\<close>
+typ "'a :: t2_space"
+text\<open>The associated operation is the predicate open that determines the open sets in the T2 space.\<close>
+term "open"
+
+section \<open>Definitions and Setup for The Relativization\<close>
+
+text\<open>We gradually define relativization of topological spaces, t2 spaces, compact and closed
+ predicates and prove that they are indeed the relativization of the original predicates.\<close>
+
+definition topological_space_on_with :: "'a set \<Rightarrow> ('a set \<Rightarrow> bool) \<Rightarrow> bool"
+ where "topological_space_on_with A \<equiv> \<lambda>open. open A \<and>
+ (\<forall>S \<subseteq> A. \<forall>T \<subseteq> A. open S \<longrightarrow> open T \<longrightarrow> open (S \<inter> T))
+ \<and> (\<forall>K \<subseteq> Pow A. (\<forall>S\<in>K. open S) \<longrightarrow> open (\<Union>K))"
+
+lemma topological_space_transfer[transfer_rule]:
+ includes lifting_syntax
+ assumes [transfer_rule]: "right_total T" "bi_unique T"
+ shows "((rel_set T ===> op=) ===> op=) (topological_space_on_with (Collect (Domainp T)))
+ class.topological_space"
+ unfolding topological_space_on_with_def[abs_def] class.topological_space_def[abs_def]
+ apply transfer_prover_start
+ apply transfer_step+
+ unfolding Pow_def Ball_Collect[symmetric]
+ by blast
+
+definition t2_space_on_with :: "'a set \<Rightarrow> ('a set \<Rightarrow> bool) \<Rightarrow> bool"
+ where "t2_space_on_with A \<equiv> \<lambda>open. topological_space_on_with A open \<and>
+ (\<forall>x \<in> A. \<forall>y \<in> A. x \<noteq> y \<longrightarrow> (\<exists>U\<subseteq>A. \<exists>V\<subseteq>A. open U \<and> open V \<and> x \<in> U \<and> y \<in> V \<and> U \<inter> V = {}))"
+
+lemma t2_space_transfer[transfer_rule]:
+ includes lifting_syntax
+ assumes [transfer_rule]: "right_total T" "bi_unique T"
+ shows "((rel_set T ===> op=) ===> op=) (t2_space_on_with (Collect (Domainp T))) class.t2_space"
+ unfolding t2_space_on_with_def[abs_def] class.t2_space_def[abs_def]
+ class.t2_space_axioms_def[abs_def]
+ apply transfer_prover_start
+ apply transfer_step+
+ unfolding Ball_Collect[symmetric]
+ by blast
+
+definition closed_with :: "('a set \<Rightarrow> bool) \<Rightarrow> 'a set \<Rightarrow> bool"
+ where "closed_with \<equiv> \<lambda>open S. open (- S)"
+
+lemma closed_closed_with: "closed s = closed_with open s"
+ unfolding closed_with_def closed_def[abs_def] ..
+
+definition closed_on_with :: "'a set \<Rightarrow> ('a set \<Rightarrow> bool) \<Rightarrow> 'a set \<Rightarrow> bool"
+ where "closed_on_with A \<equiv> \<lambda>open S. open (-S \<inter> A)"
+
+lemma closed_with_transfer[transfer_rule]:
+ includes lifting_syntax
+ assumes [transfer_rule]: "right_total T" "bi_unique T"
+ shows "((rel_set T ===> op=) ===> rel_set T===> op=) (closed_on_with (Collect (Domainp T)))
+ closed_with"
+ unfolding closed_with_def closed_on_with_def by transfer_prover
+
+definition compact_with :: "('a set \<Rightarrow> bool) \<Rightarrow> 'a set \<Rightarrow> bool"
+ where "compact_with \<equiv> \<lambda>open S. (\<forall>C. (\<forall>c\<in>C. open c) \<and> S \<subseteq> \<Union>C \<longrightarrow> (\<exists>D\<subseteq>C. finite D \<and> S \<subseteq> \<Union>D))"
+
+lemma compact_compact_with: "compact s = compact_with open s"
+ unfolding compact_with_def compact_eq_heine_borel[abs_def] ..
+
+definition compact_on_with :: "'a set \<Rightarrow> ('a set \<Rightarrow> bool) \<Rightarrow> 'a set \<Rightarrow> bool"
+ where "compact_on_with A \<equiv> \<lambda>open S. (\<forall>C\<subseteq>Pow A. (\<forall>c\<in>C. open c) \<and> S \<subseteq> \<Union>C \<longrightarrow>
+ (\<exists>D\<subseteq>C. finite D \<and> S \<subseteq> \<Union>D))"
+
+lemma compact_on_with_subset_trans: "(\<forall>C\<subseteq>Pow A. (\<forall>c\<in>C. open' c) \<and> S \<subseteq> \<Union>C \<longrightarrow>
+ (\<exists>D\<subseteq>C. finite D \<and> S \<subseteq> \<Union>D)) =
+ ((\<forall>C\<subseteq>Pow A. (\<forall>c\<in>C. open' c) \<and> S \<subseteq> \<Union>C \<longrightarrow> (\<exists>D\<subseteq>Pow A. D\<subseteq>C \<and> finite D \<and> S \<subseteq> \<Union>D)))"
+ by (meson subset_trans)
+
+lemma compact_with_transfer[transfer_rule]:
+ includes lifting_syntax
+ assumes [transfer_rule]: "right_total T" "bi_unique T"
+ shows "((rel_set T ===> op=) ===> rel_set T===> op=) (compact_on_with (Collect (Domainp T)))
+ compact_with"
+ unfolding compact_with_def compact_on_with_def
+ apply transfer_prover_start
+ apply transfer_step+
+ unfolding compact_on_with_subset_trans
+ unfolding Pow_def Ball_Collect[symmetric] Ball_def Bex_def mem_Collect_eq
+ by blast
+
+setup \<open>Sign.add_const_constraint
+ (@{const_name "open"}, SOME @{typ "'a set \<Rightarrow> bool"})\<close>
+
+text\<open>The aforementioned development can be automated. The main part is already automated
+ by the transfer_prover.\<close>
+
+section \<open>The Relativization to The Set-Based Theorem\<close>
+
+text\<open>The first step of the dictionary construction.\<close>
+lemmas dictionary_first_step = compact_imp_closed[unfolded compact_compact_with closed_closed_with]
+thm dictionary_first_step
+
+text\<open>Internalization of the type class t2_space.\<close>
+lemmas internalized_sort = dictionary_first_step[internalize_sort "'a::t2_space"]
+thm internalized_sort
+
+text\<open>We unoverload the overloaded constant open and thus finish compiling out of it.\<close>
+lemmas dictionary_second_step = internalized_sort[unoverload "open :: 'a set \<Rightarrow> bool"]
+text\<open>The theorem with internalized type classes and compiled out operations is the starting point
+ for the original relativization algorithm.\<close>
+thm dictionary_second_step
+
+
+text\<open>This is the set-based variant of the theorem compact_imp_closed.\<close>
+lemma compact_imp_closed_set_based:
+ assumes "(A::'a set) \<noteq> {}"
+ shows "\<forall>open. t2_space_on_with A open \<longrightarrow> (\<forall>S\<subseteq>A. compact_on_with A open S \<longrightarrow>
+ closed_on_with A open S)"
+proof -
+ {
+ text\<open>We define the type 'b to be isomorphic to A.\<close>
+ assume T: "\<exists>(Rep :: 'b \<Rightarrow> 'a) Abs. type_definition Rep Abs A"
+ from T obtain rep :: "'b \<Rightarrow> 'a" and abs :: "'a \<Rightarrow> 'b" where t: "type_definition rep abs A"
+ by auto
+
+ text\<open>Setup for the Transfer tool.\<close>
+ define cr_b where "cr_b == \<lambda>r a. r = rep a"
+ note type_definition_Domainp[OF t cr_b_def, transfer_domain_rule]
+ note typedef_right_total[OF t cr_b_def, transfer_rule]
+ note typedef_bi_unique[OF t cr_b_def, transfer_rule]
+
+ have ?thesis
+ text\<open>Relativization by the Transfer tool.\<close>
+ using dictionary_second_step[where 'a = 'b, untransferred, simplified]
+ by blast
+
+ } note * = this[cancel_type_definition, OF assms]
+
+ show ?thesis by (rule *)
+qed
+
+setup \<open>Sign.add_const_constraint
+ (@{const_name "open"}, SOME @{typ "'a::topological_space set \<Rightarrow> bool"})\<close>
+
+text\<open>The Final Result. We can compare the type-based and the set-based statement.\<close>
+thm compact_imp_closed compact_imp_closed_set_based
+
+declare [[show_sorts]]
+text\<open>The Final Result. This time with explicitly shown type-class annotations.\<close>
+thm compact_imp_closed compact_imp_closed_set_based
+
+end
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Types_To_Sets/Types_To_Sets.thy Wed Jan 04 14:26:19 2017 +0100
@@ -0,0 +1,27 @@
+(* Title: HOL/Types_To_Sets/Types_To_Sets.thy
+ Author: Ondřej Kunčar, TU München
+*)
+
+section \<open>From Types to Sets\<close>
+
+text \<open>This theory extends Isabelle/HOL's logic by two new inference rules
+ to allow translation of types to sets as described in
+ O. Kunčar, A. Popescu: From Types to Sets by Local Type Definitions in Higher-Order Logic
+ available at http://www21.in.tum.de/~kuncar/documents/kuncar-popescu-t2s2016-extended.pdf.\<close>
+
+theory Types_To_Sets
+ imports Main
+begin
+
+subsection \<open>Rules\<close>
+
+text\<open>The following file implements the Local Typedef Rule (LT) and extends the logic by the rule.\<close>
+ML_file "local_typedef.ML"
+
+text\<open>The following file implements the Unoverloading Rule (UO) and extends the logic by the rule.\<close>
+ML_file "unoverloading.ML"
+
+text\<open>The following file implements a derived rule that internalizes type class annotations.\<close>
+ML_file "internalize_sort.ML"
+
+end
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Types_To_Sets/internalize_sort.ML Wed Jan 04 14:26:19 2017 +0100
@@ -0,0 +1,68 @@
+(* Title: HOL/Types_To_Sets/internalize_sort.ML
+ Author: Ondřej Kunčar, TU München
+
+A derived rule (by using Thm.unconstrainT) that internalizes
+type class annotations.
+*)
+
+
+(*
+ \<phi>['a::{c_1, ..., c_n} / 'a]
+---------------------------------------------------------------------
+ class.c_1 ops_1 \<Longrightarrow> ... \<Longrightarrow> class.c_n ops_n \<Longrightarrow> \<phi>['a::type / 'a]
+
+where class.c is the locale predicate associated with type class c and
+ops are operations associated with type class c. For example:
+If c = semigroup_add, then ops = op-, op+, 0, uminus.
+If c = finite, then ops = TYPE('a::type).
+*)
+
+signature INTERNALIZE_SORT =
+sig
+ val internalize_sort: ctyp -> thm -> typ * thm
+ val internalize_sort_attr: typ -> attribute
+end;
+
+structure Internalize_Sort : INTERNALIZE_SORT =
+struct
+
+fun internalize_sort ctvar thm =
+ let
+ val thy = Thm.theory_of_thm thm;
+ val tvar = Thm.typ_of ctvar;
+
+ val ((_, assms, classes),_) = Logic.unconstrainT [] (Thm.prop_of thm);
+
+ fun is_proper_class thy = can (Axclass.get_info thy); (* trick by FH *)
+ fun reduce_to_non_proper_sort (TVar (name, sort)) =
+ TVar (name, Sign.minimize_sort thy (filter_out (is_proper_class thy) (Sign.complete_sort thy sort)))
+ | reduce_to_non_proper_sort _ = error "not supported";
+
+ val data = (map fst classes) ~~ assms;
+
+ val new_tvar = get_first (fn (tvar', ((ren_tvar, _), _)) => if tvar = tvar'
+ then SOME (reduce_to_non_proper_sort ren_tvar) else NONE) data
+ |> the_default tvar;
+
+ fun localify class = Class.rules thy class |> snd |> Thm.transfer thy;
+
+ fun discharge_of_class tvar class = Thm.of_class (Thm.global_ctyp_of thy tvar, class);
+
+ val rules = map (fn (tvar', ((ren_tvar, class), _)) => if tvar = tvar'
+ then (if Class.is_class thy class then localify class else discharge_of_class ren_tvar class)
+ else discharge_of_class ren_tvar class) data;
+ in
+ (new_tvar, (Thm.unconstrainT (Thm.strip_shyps thm) OF rules) |> Drule.zero_var_indexes)
+ end;
+
+val tvar = Args.context -- Args.typ >> (fn (_, v as TFree _) => Logic.varifyT_global v
+ | (ctxt, t) => error ("Not a type variable: " ^ Syntax.string_of_typ ctxt t));
+
+fun internalize_sort_attr tvar =
+ Thm.rule_attribute [] (fn context => fn thm =>
+ (snd (internalize_sort (Thm.ctyp_of (Context.proof_of context) tvar) thm)));
+
+val _ = Context.>> (Context.map_theory (Attrib.setup @{binding internalize_sort}
+ (tvar >> internalize_sort_attr) "internalize a sort"));
+
+end;
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Types_To_Sets/local_typedef.ML Wed Jan 04 14:26:19 2017 +0100
@@ -0,0 +1,84 @@
+(* Title: HOL/Types_To_Sets/local_typedef.ML
+ Author: Ondřej Kunčar, TU München
+
+The Local Typedef Rule (extension of the logic).
+*)
+
+signature LOCAL_TYPEDEF =
+sig
+ val cancel_type_definition: thm -> thm
+ val cancel_type_definition_attr: attribute
+end;
+
+structure Local_Typedef : LOCAL_TYPEDEF =
+struct
+
+(*
+Local Typedef Rule (LT)
+
+ \<Gamma> \<turnstile> (\<exists>(Rep::\<beta> \<Rightarrow> \<tau>) Abs. type_definition Rep Abs A) \<Longrightarrow> \<phi>
+------------------------------------------------------------- [\<beta> not in \<phi>, \<Gamma>, A;
+ \<Gamma> \<turnstile> A \<noteq> \<emptyset> \<Longrightarrow> \<phi> sort(\<beta>)=HOL.type]
+*)
+
+(** BEGINNING OF THE CRITICAL CODE **)
+
+fun dest_typedef (Const (@{const_name Ex}, _) $ Abs (_, _,
+ (Const (@{const_name Ex}, _) $ Abs (_, Abs_type,
+ (Const (@{const_name type_definition}, _)) $ Bound 1 $ Bound 0 $ set)))) =
+ (Abs_type, set)
+ | dest_typedef t = raise TERM ("dest_typedef", [t]);
+
+fun cancel_type_definition_cterm thm =
+ let
+ fun err msg = raise THM ("cancel_type_definition: " ^ msg, 0, [thm]);
+
+ val thy = Thm.theory_of_thm thm;
+ val prop = Thm.prop_of thm;
+ val hyps = Thm.hyps_of thm;
+
+ val _ = null (Thm.tpairs_of thm) orelse err "the theorem contains unsolved flex-flex pairs";
+
+ val (typedef_assm, phi) = Logic.dest_implies prop
+ handle TERM _ => err "the theorem is not an implication";
+ val (abs_type, set) = typedef_assm |> HOLogic.dest_Trueprop |> dest_typedef
+ handle TERM _ => err ("the assumption is not of the form "
+ ^ quote "\<exists>Rep Abs. type_definition Rep Abs A");
+
+ val (repT, absT) = Term.dest_funT abs_type;
+ val _ = Term.is_TVar absT orelse err "the abstract type is not a schematic type variable";
+ val (absT_name, sorts) = Term.dest_TVar absT;
+
+ val typeS = Sign.defaultS thy;
+ val _ = sorts = typeS orelse err ("the type " ^ quote (fst absT_name) ^ " is not of the sort "
+ ^ quote (Syntax.string_of_sort_global thy typeS));
+
+ fun contains_absT tm = member (op=) (Term.add_tvars tm []) (absT_name, sorts);
+ fun err_contains_absT_in msg = err (msg ^ " contains the forbidden type "
+ ^ quote (fst absT_name));
+ val _ = not (contains_absT phi) orelse err_contains_absT_in "the conclusion";
+ val _ = not (contains_absT set) orelse err_contains_absT_in "the set term";
+ (* the following test is superfluous; the meta hypotheses cannot currently contain TVars *)
+ val _ = not (exists contains_absT hyps) orelse err_contains_absT_in "one of the hypotheses";
+
+ val not_empty_assm = HOLogic.mk_Trueprop
+ (HOLogic.mk_not (HOLogic.mk_eq (set, HOLogic.mk_set repT [])));
+ val prop = Logic.list_implies (hyps @ [not_empty_assm], phi);
+ in
+ Thm.global_cterm_of thy prop |> Thm.weaken_sorts (Thm.shyps_of thm)
+ end;
+
+(** END OF THE CRITICAL CODE **)
+
+val (_, cancel_type_definition_oracle) = Context.>>> (Context.map_theory_result
+ (Thm.add_oracle (@{binding cancel_type_definition}, cancel_type_definition_cterm)));
+
+fun cancel_type_definition thm =
+ Drule.implies_elim_list (cancel_type_definition_oracle thm) (map Thm.assume (Thm.chyps_of thm));
+
+val cancel_type_definition_attr = Thm.rule_attribute [] (K cancel_type_definition);
+
+val _ = Context.>> (Context.map_theory (Attrib.setup @{binding cancel_type_definition}
+ (Scan.succeed cancel_type_definition_attr) "cancel a local type definition"));
+
+end;
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Types_To_Sets/unoverloading.ML Wed Jan 04 14:26:19 2017 +0100
@@ -0,0 +1,140 @@
+(* Title: HOL/Types_To_Sets/unoverloading.ML
+ Author: Ondřej Kunčar, TU München
+
+The Unoverloading Rule (extension of the logic).
+*)
+
+signature UNOVERLOADING =
+sig
+ val unoverload: cterm -> thm -> thm
+ val unoverload_attr: cterm -> attribute
+end;
+
+structure Unoverloading : UNOVERLOADING =
+struct
+
+(*
+Unoverloading Rule (UO)
+
+ \<turnstile> \<phi>[c::\<sigma> / x::\<sigma>]
+---------------------------- [no type or constant or type class in \<phi>
+ \<turnstile> \<And>x::\<sigma>. \<phi> depends on c::\<sigma>; c::\<sigma> is undefined]
+*)
+
+(* The following functions match_args, reduction and entries_of were
+ cloned from defs.ML and theory.ML because the functions are not public.
+ Notice that these functions already belong to the critical code.
+*)
+
+(* >= *)
+fun match_args (Ts, Us) =
+ if Type.could_matches (Ts, Us) then
+ Option.map Envir.subst_type
+ (SOME (Type.raw_matches (Ts, Us) Vartab.empty) handle Type.TYPE_MATCH => NONE)
+ else NONE;
+
+fun reduction defs (deps : Defs.entry list) : Defs.entry list option =
+ let
+ fun reduct Us (Ts, rhs) =
+ (case match_args (Ts, Us) of
+ NONE => NONE
+ | SOME subst => SOME (map (apsnd (map subst)) rhs));
+ fun reducts (d, Us) = get_first (reduct Us) (Defs.get_deps defs d);
+
+ val reds = map (`reducts) deps;
+ val deps' =
+ if forall (is_none o #1) reds then NONE
+ else SOME (fold_rev
+ (fn (NONE, dp) => insert (op =) dp | (SOME dps, _) => fold (insert (op =)) dps) reds []);
+ in deps' end;
+
+fun const_and_typ_entries_of thy tm =
+ let
+ val consts =
+ fold_aterms (fn Const const => insert (op =) (Theory.const_dep thy const) | _ => I) tm [];
+ val types =
+ (fold_types o fold_subtypes) (fn Type t => insert (op =) (Theory.type_dep t) | _ => I) tm [];
+ in
+ consts @ types
+ end;
+
+
+(* The actual implementation *)
+
+(** BEGINNING OF THE CRITICAL CODE **)
+
+fun fold_atyps_classes f =
+ fold_atyps (fn T as TFree (_, S) => fold (pair T #> f) S
+ | T as TVar (_, S) => fold (pair T #> f) S
+ (* just to avoid a warning about incomplete patterns *)
+ | _ => raise TERM ("fold_atyps_classes", []));
+
+fun class_entries_of thy tm =
+ let
+ val var_classes = (fold_types o fold_atyps_classes) (insert op=) tm [];
+ in
+ map (Logic.mk_of_class #> Term.head_of #> Term.dest_Const #> Theory.const_dep thy) var_classes
+ end;
+
+fun unoverload_impl cconst thm =
+ let
+ fun err msg = raise THM ("unoverloading: " ^ msg, 0, [thm]);
+
+ val _ = null (Thm.hyps_of thm) orelse err "the theorem has meta hypotheses";
+ val _ = null (Thm.tpairs_of thm) orelse err "the theorem contains unresolved flex-flex pairs";
+
+ val prop = Thm.prop_of thm;
+ val prop_tfrees = rev (Term.add_tfree_names prop []);
+ val _ = null prop_tfrees orelse err ("the theorem contains free type variables "
+ ^ commas_quote prop_tfrees);
+
+ val const = Thm.term_of cconst;
+ val _ = Term.is_Const const orelse err "the parameter is is not a constant";
+ val const_tfrees = rev (Term.add_tfree_names const []);
+ val _ = null const_tfrees orelse err ("the constant contains free type variables "
+ ^ commas_quote const_tfrees);
+
+ val thy = Thm.theory_of_thm thm;
+ val defs = Theory.defs_of thy;
+
+ val const_entry = Theory.const_dep thy (Term.dest_Const const);
+
+ val Uss = Defs.specifications_of defs (fst const_entry);
+ val _ = forall (fn spec => is_none (match_args (#lhs spec, snd const_entry))) Uss
+ orelse err "the constant instance has already a specification";
+
+ val context = Defs.global_context thy;
+ val prt_entry = Pretty.string_of o Defs.pretty_entry context;
+
+ fun dep_err (c, Ts) (d, Us) =
+ err (prt_entry (c, Ts) ^ " depends on " ^ prt_entry (d, Us));
+ fun deps_of entry = perhaps_loop (reduction defs) [entry] |> these;
+ fun not_depends_on_const prop_entry = forall (not_equal const_entry) (deps_of prop_entry)
+ orelse dep_err prop_entry const_entry;
+ val prop_entries = const_and_typ_entries_of thy prop @ class_entries_of thy prop;
+ val _ = forall not_depends_on_const prop_entries;
+ in
+ Thm.global_cterm_of thy (Logic.all const prop) |> Thm.weaken_sorts (Thm.shyps_of thm)
+ end;
+
+(** END OF THE CRITICAL CODE **)
+
+val (_, unoverload_oracle) = Context.>>> (Context.map_theory_result
+ (Thm.add_oracle (@{binding unoverload},
+ fn (const, thm) => unoverload_impl const thm)));
+
+fun unoverload const thm = unoverload_oracle (const, thm);
+
+fun unoverload_attr const =
+ Thm.rule_attribute [] (fn _ => fn thm => (unoverload const thm));
+
+val const = Args.context -- Args.term >>
+ (fn (ctxt, tm) =>
+ if not (Term.is_Const tm)
+ then error ("The term is not a constant: " ^ Syntax.string_of_term ctxt tm)
+ else tm |> Logic.varify_types_global |> Thm.cterm_of ctxt);
+
+val _ = Context.>> (Context.map_theory (Attrib.setup @{binding unoverload}
+ (const >> unoverload_attr) "unoverload an uninterpreted constant"));
+
+end;
--- a/src/HOL/Wellfounded.thy Wed Jan 04 14:26:08 2017 +0100
+++ b/src/HOL/Wellfounded.thy Wed Jan 04 14:26:19 2017 +0100
@@ -308,7 +308,7 @@
done
lemma wfP_SUP:
- "\<forall>i. wfP (r i) \<Longrightarrow> \<forall>i j. r i \<noteq> r j \<longrightarrow> inf (DomainP (r i)) (RangeP (r j)) = bot \<Longrightarrow>
+ "\<forall>i. wfP (r i) \<Longrightarrow> \<forall>i j. r i \<noteq> r j \<longrightarrow> inf (Domainp (r i)) (Rangep (r j)) = bot \<Longrightarrow>
wfP (SUPREMUM UNIV r)"
by (rule wf_UN[to_pred]) simp_all
--- a/src/HOL/Word/Bit_Representation.thy Wed Jan 04 14:26:08 2017 +0100
+++ b/src/HOL/Word/Bit_Representation.thy Wed Jan 04 14:26:19 2017 +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 Wed Jan 04 14:26:08 2017 +0100
+++ b/src/HOL/Word/Bits_Int.thy Wed Jan 04 14:26:19 2017 +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 Wed Jan 04 14:26:08 2017 +0100
+++ b/src/HOL/Word/Misc_Numeric.thy Wed Jan 04 14:26:19 2017 +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 Wed Jan 04 14:26:08 2017 +0100
+++ b/src/HOL/Word/Word.thy Wed Jan 04 14:26:19 2017 +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 Wed Jan 04 14:26:08 2017 +0100
+++ b/src/HOL/Word/Word_Miscellaneous.thy Wed Jan 04 14:26:19 2017 +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 Wed Jan 04 14:26:08 2017 +0100
+++ b/src/HOL/ex/Word_Type.thy Wed Jan 04 14:26:19 2017 +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>
--- a/src/Provers/clasimp.ML Wed Jan 04 14:26:08 2017 +0100
+++ b/src/Provers/clasimp.ML Wed Jan 04 14:26:19 2017 +0100
@@ -193,9 +193,9 @@
val iffN = "iff";
val iff_modifiers =
- [Args.$$$ iffN -- Scan.option Args.add -- Args.colon >> K (Method.modifier iff_add @{here}),
- Args.$$$ iffN -- Scan.option Args.add -- Args.query_colon >> K (Method.modifier iff_add' @{here}),
- Args.$$$ iffN -- Args.del -- Args.colon >> K (Method.modifier iff_del @{here})];
+ [Args.$$$ iffN -- Scan.option Args.add -- Args.colon >> K (Method.modifier iff_add \<^here>),
+ Args.$$$ iffN -- Scan.option Args.add -- Args.query_colon >> K (Method.modifier iff_add' \<^here>),
+ Args.$$$ iffN -- Args.del -- Args.colon >> K (Method.modifier iff_del \<^here>)];
val clasimp_modifiers =
Simplifier.simp_modifiers @ Splitter.split_modifiers @
--- a/src/Provers/classical.ML Wed Jan 04 14:26:08 2017 +0100
+++ b/src/Provers/classical.ML Wed Jan 04 14:26:19 2017 +0100
@@ -947,13 +947,13 @@
(* automatic methods *)
val cla_modifiers =
- [Args.$$$ destN -- Args.bang_colon >> K (Method.modifier (safe_dest NONE) @{here}),
- Args.$$$ destN -- Args.colon >> K (Method.modifier (unsafe_dest NONE) @{here}),
- Args.$$$ elimN -- Args.bang_colon >> K (Method.modifier (safe_elim NONE) @{here}),
- Args.$$$ elimN -- Args.colon >> K (Method.modifier (unsafe_elim NONE) @{here}),
- Args.$$$ introN -- Args.bang_colon >> K (Method.modifier (safe_intro NONE) @{here}),
- Args.$$$ introN -- Args.colon >> K (Method.modifier (unsafe_intro NONE) @{here}),
- Args.del -- Args.colon >> K (Method.modifier rule_del @{here})];
+ [Args.$$$ destN -- Args.bang_colon >> K (Method.modifier (safe_dest NONE) \<^here>),
+ Args.$$$ destN -- Args.colon >> K (Method.modifier (unsafe_dest NONE) \<^here>),
+ Args.$$$ elimN -- Args.bang_colon >> K (Method.modifier (safe_elim NONE) \<^here>),
+ Args.$$$ elimN -- Args.colon >> K (Method.modifier (unsafe_elim NONE) \<^here>),
+ Args.$$$ introN -- Args.bang_colon >> K (Method.modifier (safe_intro NONE) \<^here>),
+ Args.$$$ introN -- Args.colon >> K (Method.modifier (unsafe_intro NONE) \<^here>),
+ Args.del -- Args.colon >> K (Method.modifier rule_del \<^here>)];
fun cla_method tac = Method.sections cla_modifiers >> K (SIMPLE_METHOD o tac);
fun cla_method' tac = Method.sections cla_modifiers >> K (SIMPLE_METHOD' o tac);
--- a/src/Provers/splitter.ML Wed Jan 04 14:26:08 2017 +0100
+++ b/src/Provers/splitter.ML Wed Jan 04 14:26:19 2017 +0100
@@ -473,9 +473,9 @@
(* methods *)
val split_modifiers =
- [Args.$$$ splitN -- Args.colon >> K (Method.modifier (split_add false) @{here}),
- Args.$$$ splitN -- Args.bang_colon >> K (Method.modifier (split_add true) @{here}),
- Args.$$$ splitN -- Args.del -- Args.colon >> K (Method.modifier split_del @{here})];
+ [Args.$$$ splitN -- Args.colon >> K (Method.modifier (split_add false) \<^here>),
+ Args.$$$ splitN -- Args.bang_colon >> K (Method.modifier (split_add true) \<^here>),
+ Args.$$$ splitN -- Args.del -- Args.colon >> K (Method.modifier split_del \<^here>)];
val _ =
Theory.setup
--- a/src/Pure/Admin/check_sources.scala Wed Jan 04 14:26:08 2017 +0100
+++ b/src/Pure/Admin/check_sources.scala Wed Jan 04 14:26:19 2017 +0100
@@ -25,9 +25,9 @@
try {
Symbol.decode_strict(line)
- for { c <- Word.codepoint_iterator(line); if c > 128 && !Character.isAlphabetic(c) }
+ for { c <- Codepoint.iterator(line); if c > 128 && !Character.isAlphabetic(c) }
{
- Output.warning("Suspicious Unicode character " + quote(Word.codepoint(c)) +
+ Output.warning("Suspicious Unicode character " + quote(Codepoint.string(c)) +
Position.here(line_pos(i)))
}
}
--- a/src/Pure/Admin/ci_api.scala Wed Jan 04 14:26:08 2017 +0100
+++ b/src/Pure/Admin/ci_api.scala Wed Jan 04 14:26:19 2017 +0100
@@ -32,7 +32,7 @@
def build_jobs(): List[String] =
for {
- job <- JSON.array(invoke(root()), "jobs")
+ job <- JSON.array(invoke(root()), "jobs").getOrElse(Nil)
_class <- JSON.string(job, "_class")
if _class == "hudson.model.FreeStyleProject"
name <- JSON.string(job, "name")
@@ -56,7 +56,8 @@
for {
build <- JSON.array(
- invoke(root() + "/job/" + job_name, "tree=builds[number,timestamp,artifacts[*]]"), "builds")
+ invoke(root() + "/job/" + job_name, "tree=builds[number,timestamp,artifacts[*]]"), "builds")
+ .getOrElse(Nil)
number <- JSON.int(build, "number")
timestamp <- JSON.long(build, "timestamp")
} yield {
@@ -64,7 +65,7 @@
val output = Url(job_prefix + "/consoleText")
val session_logs =
for {
- artifact <- JSON.array(build, "artifacts")
+ artifact <- JSON.array(build, "artifacts").getOrElse(Nil)
log_path <- JSON.string(artifact, "relativePath")
session <- (log_path match { case Log_Session(name) => Some(name) case _ => None })
} yield (session -> Url(job_prefix + "/artifact/" + log_path))
--- a/src/Pure/Admin/isabelle_cronjob.scala Wed Jan 04 14:26:08 2017 +0100
+++ b/src/Pure/Admin/isabelle_cronjob.scala Wed Jan 04 14:26:19 2017 +0100
@@ -97,7 +97,7 @@
private val remote_builds =
List(
List(Remote_Build("lxbroy8",
- options = "-m32 -B -M1x2,2 -t polyml-test -e 'init_component /home/isabelle/contrib/polyml-test-7a7b742897e9'",
+ options = "-m32 -B -M1x2,2 -t polyml-test -e 'init_component /home/isabelle/contrib/polyml-test-8529546198aa'",
args = "-N -g timing")),
List(Remote_Build("lxbroy9", options = "-m32 -B -M1x2,2", args = "-N -g timing")),
List(Remote_Build("lxbroy10", options = "-m32 -B -M1x4,2,4,6", args = "-N -g timing")),
--- a/src/Pure/Concurrent/multithreading.ML Wed Jan 04 14:26:08 2017 +0100
+++ b/src/Pure/Concurrent/multithreading.ML Wed Jan 04 14:26:19 2017 +0100
@@ -48,7 +48,7 @@
fun sync_wait time cond lock =
Thread_Attributes.with_attributes
- (Thread_Attributes.sync_interrupts (Thread.getAttributes ()))
+ (Thread_Attributes.sync_interrupts (Thread_Attributes.get_attributes ()))
(fn _ =>
(case time of
SOME t => Exn.Res (ConditionVar.waitUntil (cond, lock, t))
@@ -61,7 +61,7 @@
val trace = ref 0;
fun tracing level msg =
- if level > ! trace then ()
+ if ! trace < level then ()
else Thread_Attributes.uninterruptible (fn _ => fn () =>
(TextIO.output (TextIO.stdErr, (">>> " ^ msg () ^ "\n")); TextIO.flushOut TextIO.stdErr)
handle _ (*sic*) => ()) ();
@@ -79,20 +79,27 @@
fun synchronized name lock e =
Exn.release (Thread_Attributes.uninterruptible (fn restore_attributes => fn () =>
- let
- val immediate =
- if Mutex.trylock lock then true
- else
- let
- val _ = tracing 5 (fn () => name ^ ": locking ...");
- val timer = Timer.startRealTimer ();
- val _ = Mutex.lock lock;
- val time = Timer.checkRealTimer timer;
- val _ = tracing_time true time (fn () => name ^ ": locked after " ^ Time.toString time);
- in false end;
- val result = Exn.capture (restore_attributes e) ();
- val _ = if immediate then () else tracing 5 (fn () => name ^ ": unlocking ...");
- val _ = Mutex.unlock lock;
- in result end) ());
+ if ! trace > 0 then
+ let
+ val immediate =
+ if Mutex.trylock lock then true
+ else
+ let
+ val _ = tracing 5 (fn () => name ^ ": locking ...");
+ val timer = Timer.startRealTimer ();
+ val _ = Mutex.lock lock;
+ val time = Timer.checkRealTimer timer;
+ val _ = tracing_time true time (fn () => name ^ ": locked after " ^ Time.toString time);
+ in false end;
+ val result = Exn.capture (restore_attributes e) ();
+ val _ = if immediate then () else tracing 5 (fn () => name ^ ": unlocking ...");
+ val _ = Mutex.unlock lock;
+ in result end
+ else
+ let
+ val _ = Mutex.lock lock;
+ val result = Exn.capture (restore_attributes e) ();
+ val _ = Mutex.unlock lock;
+ in result end) ());
end;
--- a/src/Pure/Concurrent/standard_thread.ML Wed Jan 04 14:26:08 2017 +0100
+++ b/src/Pure/Concurrent/standard_thread.ML Wed Jan 04 14:26:19 2017 +0100
@@ -50,7 +50,8 @@
fun attributes ({stack_limit, interrupts, ...}: params) =
Thread.MaximumMLStack stack_limit ::
- (if interrupts then Thread_Attributes.public_interrupts else Thread_Attributes.no_interrupts);
+ Thread_Attributes.convert_attributes
+ (if interrupts then Thread_Attributes.public_interrupts else Thread_Attributes.no_interrupts);
fun fork (params: params) body =
Thread.fork (fn () =>
--- a/src/Pure/Concurrent/thread_attributes.ML Wed Jan 04 14:26:08 2017 +0100
+++ b/src/Pure/Concurrent/thread_attributes.ML Wed Jan 04 14:26:19 2017 +0100
@@ -6,7 +6,10 @@
signature THREAD_ATTRIBUTES =
sig
- type attributes = Thread.Thread.threadAttribute list
+ type attributes
+ val get_attributes: unit -> attributes
+ val set_attributes: attributes -> unit
+ val convert_attributes: attributes -> Thread.Thread.threadAttribute list
val no_interrupts: attributes
val test_interrupts: attributes
val public_interrupts: attributes
@@ -21,41 +24,84 @@
structure Thread_Attributes: THREAD_ATTRIBUTES =
struct
-type attributes = Thread.Thread.threadAttribute list;
+abstype attributes = Attributes of Word.word
+with
-val no_interrupts =
- [Thread.Thread.EnableBroadcastInterrupt false,
- Thread.Thread.InterruptState Thread.Thread.InterruptDefer];
+(* thread attributes *)
+
+val thread_attributes = 0w7;
+
+val broadcast = 0w1;
-val test_interrupts =
- [Thread.Thread.EnableBroadcastInterrupt false,
- Thread.Thread.InterruptState Thread.Thread.InterruptSynch];
+val interrupt_state = 0w6;
+val interrupt_defer = 0w0;
+val interrupt_synch = 0w2;
+val interrupt_asynch = 0w4;
+val interrupt_asynch_once = 0w6;
+
-val public_interrupts =
- [Thread.Thread.EnableBroadcastInterrupt true,
- Thread.Thread.InterruptState Thread.Thread.InterruptAsynchOnce];
+(* access thread flags *)
+
+val thread_flags = 0w1;
+
+fun load_word () : word =
+ RunCall.loadWord (RunCall.unsafeCast (Thread.Thread.self ()), thread_flags);
+
+fun get_attributes () = Attributes (Word.andb (load_word (), thread_attributes));
-val private_interrupts =
- [Thread.Thread.EnableBroadcastInterrupt false,
- Thread.Thread.InterruptState Thread.Thread.InterruptAsynchOnce];
+fun set_attributes (Attributes w) =
+ let
+ val w' = Word.orb (Word.andb (load_word (), Word.notb thread_attributes), w);
+ val _ = RunCall.storeWord (RunCall.unsafeCast (Thread.Thread.self ()), thread_flags, w');
+ in
+ if Word.andb (w', interrupt_asynch) = interrupt_asynch
+ then Thread.Thread.testInterrupt () else ()
+ end;
+
+fun convert_attributes (Attributes w) =
+ [Thread.Thread.EnableBroadcastInterrupt (Word.andb (w, broadcast) = broadcast),
+ Thread.Thread.InterruptState
+ let val w' = Word.andb (w, interrupt_state) in
+ if w' = interrupt_defer then Thread.Thread.InterruptDefer
+ else if w' = interrupt_synch then Thread.Thread.InterruptSynch
+ else if w' = interrupt_asynch then Thread.Thread.InterruptAsynch
+ else Thread.Thread.InterruptAsynchOnce
+ end];
-val sync_interrupts = map
- (fn x as Thread.Thread.InterruptState Thread.Thread.InterruptDefer => x
- | Thread.Thread.InterruptState _ => Thread.Thread.InterruptState Thread.Thread.InterruptSynch
- | x => x);
+
+(* common configurations *)
+
+val no_interrupts = Attributes interrupt_defer;
+val test_interrupts = Attributes interrupt_synch;
+val public_interrupts = Attributes (Word.orb (broadcast, interrupt_asynch_once));
+val private_interrupts = Attributes interrupt_asynch_once;
-val safe_interrupts = map
- (fn Thread.Thread.InterruptState Thread.Thread.InterruptAsynch =>
- Thread.Thread.InterruptState Thread.Thread.InterruptAsynchOnce
- | x => x);
+fun sync_interrupts (Attributes w) =
+ let
+ val w' = Word.andb (w, interrupt_state);
+ val w'' = Word.andb (w, Word.notb interrupt_state);
+ in Attributes (if w' = interrupt_defer then w else Word.orb (w'', interrupt_synch)) end;
+
+fun safe_interrupts (Attributes w) =
+ let
+ val w' = Word.andb (w, interrupt_state);
+ val w'' = Word.andb (w, Word.notb interrupt_state);
+ in Attributes (if w' = interrupt_asynch then Word.orb (w'', interrupt_asynch_once) else w) end;
fun with_attributes new_atts e =
let
- val orig_atts = safe_interrupts (Thread.Thread.getAttributes ());
- val result = Exn.capture (fn () =>
- (Thread.Thread.setAttributes (safe_interrupts new_atts); e orig_atts)) ();
- val _ = Thread.Thread.setAttributes orig_atts;
- in Exn.release result end;
+ val atts1 = safe_interrupts (get_attributes ());
+ val atts2 = safe_interrupts new_atts;
+ in
+ if atts1 = atts2 then e atts1
+ else
+ let
+ val result = Exn.capture (fn () => (set_attributes atts2; e atts1)) ();
+ val _ = set_attributes atts1;
+ in Exn.release result end
+ end;
+
+end;
fun uninterruptible f x =
with_attributes no_interrupts (fn atts =>
@@ -63,10 +109,10 @@
fun expose_interrupt () =
let
- val orig_atts = safe_interrupts (Thread.Thread.getAttributes ());
- val _ = Thread.Thread.setAttributes test_interrupts;
+ val orig_atts = safe_interrupts (get_attributes ());
+ val _ = set_attributes test_interrupts;
val test = Exn.capture Thread.Thread.testInterrupt ();
- val _ = Thread.Thread.setAttributes orig_atts;
+ val _ = set_attributes orig_atts;
in Exn.release test end;
end;
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/Pure/General/codepoint.scala Wed Jan 04 14:26:19 2017 +0100
@@ -0,0 +1,53 @@
+/* Title: Pure/General/codepoint.scala
+ Author: Makarius
+
+Unicode codepoints vs. Unicode string encoding.
+*/
+
+package isabelle
+
+
+object Codepoint
+{
+ def string(c: Int): String = new String(Array(c), 0, 1)
+
+ def iterator(s: String): Iterator[Int] =
+ new Iterator[Int] {
+ var offset = 0
+ def hasNext: Boolean = offset < s.length
+ def next: Int =
+ {
+ val c = s.codePointAt(offset)
+ offset += Character.charCount(c)
+ c
+ }
+ }
+
+ def length(s: String): Int = iterator(s).length
+
+ trait Length extends Text.Length
+ {
+ protected def codepoint_length(c: Int): Int = 1
+
+ def apply(s: String): Int =
+ (0 /: iterator(s)) { case (n, c) => n + codepoint_length(c) }
+
+ def offset(s: String, i: Int): Option[Text.Offset] =
+ {
+ if (i <= 0 || s.forall(_ < 0x80)) Text.Length.offset(s, i)
+ else {
+ val length = s.length
+ var offset = 0
+ var j = 0
+ while (offset < length && j < i) {
+ val c = s.codePointAt(offset)
+ offset += Character.charCount(c)
+ j += codepoint_length(c)
+ }
+ if (j >= i) Some(offset) else None
+ }
+ }
+ }
+
+ object Length extends Length
+}
--- a/src/Pure/General/file.ML Wed Jan 04 14:26:08 2017 +0100
+++ b/src/Pure/General/file.ML Wed Jan 04 14:26:19 2017 +0100
@@ -1,7 +1,7 @@
(* Title: Pure/General/file.ML
- Author: Markus Wenzel, TU Muenchen
+ Author: Makarius
-File system operations.
+File-system operations.
*)
signature FILE =
--- a/src/Pure/General/file.scala Wed Jan 04 14:26:08 2017 +0100
+++ b/src/Pure/General/file.scala Wed Jan 04 14:26:19 2017 +0100
@@ -1,7 +1,7 @@
/* Title: Pure/General/file.scala
Author: Makarius
-File system operations.
+File-system operations.
*/
package isabelle
@@ -33,7 +33,7 @@
if (Platform.is_windows) {
val Platform_Root = new Regex("(?i)" +
Pattern.quote(Isabelle_System.cygwin_root()) + """(?:\\+|\z)(.*)""")
- val Drive = new Regex("""([a-zA-Z]):\\*(.*)""")
+ val Drive = new Regex("""\\?([a-zA-Z]):\\*(.*)""")
platform_path.replace('/', '\\') match {
case Platform_Root(rest) => "/" + rest.replace('\\', '/')
@@ -48,6 +48,7 @@
def standard_path(file: JFile): String = standard_path(file.getPath)
def path(file: JFile): Path = Path.explode(standard_path(file))
+ def pwd(): Path = path(Path.current.file.toPath.toAbsolutePath.toFile)
def standard_url(name: String): String =
try {
@@ -95,16 +96,6 @@
def platform_path(path: Path): String = platform_path(standard_path(path))
def platform_file(path: Path): JFile = new JFile(platform_path(path))
- def platform_url(raw_path: Path): String =
- {
- val path = raw_path.expand
- require(path.is_absolute)
- val s = platform_path(path).replaceAll(" ", "%20")
- if (!Platform.is_windows) "file://" + s
- else if (s.startsWith("\\\\")) "file:" + s.replace('\\', '/')
- else "file:///" + s.replace('\\', '/')
- }
-
/* bash path */
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/Pure/General/file_watcher.scala Wed Jan 04 14:26:19 2017 +0100
@@ -0,0 +1,108 @@
+/* Title: Pure/General/file_watcher.scala
+ Author: Makarius
+
+Watcher for file-system events.
+*/
+
+package isabelle
+
+
+import java.io.{File => JFile}
+import java.nio.file.FileSystems
+import java.nio.file.{WatchKey, WatchEvent, Path => JPath}
+import java.nio.file.StandardWatchEventKinds.{ENTRY_CREATE, ENTRY_DELETE, ENTRY_MODIFY}
+
+import scala.collection.JavaConversions
+
+
+object File_Watcher
+{
+ def apply(handle: Set[JFile] => Unit, delay: => Time = Time.seconds(0.5)): File_Watcher =
+ new File_Watcher(handle, delay)
+
+
+ /* internal state */
+
+ sealed case class State(
+ dirs: Map[JFile, WatchKey] = Map.empty,
+ changed: Set[JFile] = Set.empty)
+}
+
+class File_Watcher private(handle: Set[JFile] => Unit, delay: Time)
+{
+ private val state = Synchronized(File_Watcher.State())
+ private val watcher = FileSystems.getDefault.newWatchService()
+
+
+ /* registered directories */
+
+ def register(dir: JFile): Unit =
+ state.change(st =>
+ st.dirs.get(dir) match {
+ case Some(key) if key.isValid => st
+ case _ =>
+ val key = dir.toPath.register(watcher, ENTRY_CREATE, ENTRY_DELETE, ENTRY_MODIFY)
+ st.copy(dirs = st.dirs + (dir -> key))
+ })
+
+ def deregister(dir: JFile): Unit =
+ state.change(st =>
+ st.dirs.get(dir) match {
+ case None => st
+ case Some(key) =>
+ key.cancel
+ st.copy(dirs = st.dirs - dir)
+ })
+
+
+ /* changed directory entries */
+
+ private val delay_changed = Standard_Thread.delay_last(delay)
+ {
+ val changed = state.change_result(st => (st.changed, st.copy(changed = Set.empty)))
+ handle(changed)
+ }
+
+ private val watcher_thread = Standard_Thread.fork("File_Watcher", daemon = true)
+ {
+ try {
+ while (true) {
+ val key = watcher.take
+ val has_changed =
+ state.change_result(st =>
+ {
+ val (remove, changed) =
+ st.dirs.collectFirst({ case (dir, key1) if key == key1 => dir }) match {
+ case Some(dir) =>
+ val events =
+ JavaConversions.collectionAsScalaIterable(
+ key.pollEvents.asInstanceOf[java.util.List[WatchEvent[JPath]]])
+ val remove = if (key.reset) None else Some(dir)
+ val changed =
+ (Set.empty[JFile] /: events.iterator) {
+ case (set, event) => set + dir.toPath.resolve(event.context).toFile
+ }
+ (remove, changed)
+ case None =>
+ key.pollEvents
+ key.reset
+ (None, Set.empty[JFile])
+ }
+ (changed.nonEmpty, st.copy(dirs = st.dirs -- remove, changed = st.changed ++ changed))
+ })
+ if (has_changed) delay_changed.invoke()
+ }
+ }
+ catch { case Exn.Interrupt() => }
+ }
+
+
+ /* shutdown */
+
+ def shutdown()
+ {
+ watcher_thread.interrupt
+ watcher_thread.join
+ delay_changed.revoke
+ }
+}
--- a/src/Pure/General/json.scala Wed Jan 04 14:26:08 2017 +0100
+++ b/src/Pure/General/json.scala Wed Jan 04 14:26:19 2017 +0100
@@ -7,48 +7,175 @@
package isabelle
+import scala.util.parsing.json.{JSONObject, JSONArray}
+
object JSON
{
- /* parse */
-
- def parse(text: String): Any =
- scala.util.parsing.json.JSON.parseFull(text) getOrElse error("Malformed JSON")
+ type T = Any
+ type S = String
- /* field access */
+ /* standard format */
+
+ def parse(s: S): T = Format.unapply(s) getOrElse error("Malformed JSON")
+
+ object Format
+ {
+ def unapply(s: S): Option[T] =
+ try { scala.util.parsing.json.JSON.parseFull(s) }
+ catch { case ERROR(_) => None }
+
+ def apply(json: T): S =
+ {
+ val result = new StringBuilder
+
+ def string(s: String)
+ {
+ result += '"'
+ result ++=
+ s.iterator.map {
+ case '"' => "\\\""
+ case '\\' => "\\\\"
+ case '\b' => "\\b"
+ case '\f' => "\\f"
+ case '\n' => "\\n"
+ case '\r' => "\\r"
+ case '\t' => "\\t"
+ case c =>
+ if (c <= '\u001f' || c >= '\u007f' && c <= '\u009f') "\\u%04x".format(c.toInt)
+ else c
+ }.mkString
+ result += '"'
+ }
+
+ def array(list: List[T])
+ {
+ result += '['
+ Library.separate(None, list.map(Some(_))).foreach({
+ case None => result += ','
+ case Some(x) => json_format(x)
+ })
+ result += ']'
+ }
+
+ def object_(obj: Map[String, T])
+ {
+ result += '{'
+ Library.separate(None, obj.toList.map(Some(_))).foreach({
+ case None => result += ','
+ case Some((x, y)) =>
+ string(x)
+ result += ':'
+ json_format(y)
+ })
+ result += '}'
+ }
- def any(obj: Any, name: String): Option[Any] =
+ def json_format(x: T)
+ {
+ x match {
+ case null => result ++= "null"
+ case _: Int | _: Long | _: Boolean => result ++= x.toString
+ case n: Double =>
+ val i = n.toLong
+ result ++= (if (i.toDouble == n) i.toString else n.toString)
+ case s: String => string(s)
+ case JSONObject(obj) => object_(obj)
+ case obj: Map[_, _] if obj.keySet.forall(_.isInstanceOf[String]) =>
+ object_(obj.asInstanceOf[Map[String, T]])
+ case JSONArray(list) => array(list)
+ case list: List[T] => array(list)
+ case _ => error("Bad JSON value: " + x.toString)
+ }
+ }
+
+ json_format(json)
+ result.toString
+ }
+ }
+
+
+ /* numbers */
+
+ object Number
+ {
+ object Double {
+ def unapply(json: T): Option[scala.Double] =
+ json match {
+ case x: scala.Double => Some(x)
+ case x: scala.Long => Some(x.toDouble)
+ case x: scala.Int => Some(x.toDouble)
+ case _ => None
+ }
+ }
+
+ object Long {
+ def unapply(json: T): Option[scala.Long] =
+ json match {
+ case x: scala.Double if x.toLong.toDouble == x => Some(x.toLong)
+ case x: scala.Long => Some(x)
+ case x: scala.Int => Some(x.toLong)
+ case _ => None
+ }
+ }
+
+ object Int {
+ def unapply(json: T): Option[scala.Int] =
+ json match {
+ case x: scala.Double if x.toInt.toDouble == x => Some(x.toInt)
+ case x: scala.Long if x.toInt.toLong == x => Some(x.toInt)
+ case x: scala.Int => Some(x)
+ case _ => None
+ }
+ }
+ }
+
+
+ /* object values */
+
+ def value(obj: T, name: String): Option[T] =
obj match {
- case m: Map[String, Any] => m.get(name)
+ case m: Map[_, _] if m.keySet.forall(_.isInstanceOf[String]) =>
+ m.asInstanceOf[Map[String, T]].get(name)
case _ => None
}
- def array(obj: Any, name: String): List[Any] =
- any(obj, name) match {
- case Some(a: List[Any]) => a
- case _ => Nil
+ def value[A](obj: T, name: String, unapply: T => Option[A]): Option[A] =
+ for {
+ json <- value(obj, name)
+ x <- unapply(json)
+ } yield x
+
+ def array(obj: T, name: String): Option[List[T]] =
+ value(obj, name) match {
+ case Some(a: List[T]) => Some(a)
+ case _ => None
}
- def string(obj: Any, name: String): Option[String] =
- any(obj, name) match {
+ def array[A](obj: T, name: String, unapply: T => Option[A]): Option[List[A]] =
+ for {
+ a0 <- array(obj, name)
+ a = a0.map(unapply(_))
+ if a.forall(_.isDefined)
+ } yield a.map(_.get)
+
+ def string(obj: T, name: String): Option[String] =
+ value(obj, name) match {
case Some(x: String) => Some(x)
case _ => None
}
- def double(obj: Any, name: String): Option[Double] =
- any(obj, name) match {
- case Some(x: Double) => Some(x)
- case _ => None
- }
+ def double(obj: T, name: String): Option[Double] =
+ value(obj, name, Number.Double.unapply)
+
+ def long(obj: T, name: String): Option[Long] =
+ value(obj, name, Number.Long.unapply)
- def long(obj: Any, name: String): Option[Long] =
- double(obj, name).map(_.toLong)
+ def int(obj: T, name: String): Option[Int] =
+ value(obj, name, Number.Int.unapply)
- def int(obj: Any, name: String): Option[Int] =
- double(obj, name).map(_.toInt)
-
- def bool(obj: Any, name: String): Option[Boolean] =
- any(obj, name) match {
+ def bool(obj: T, name: String): Option[Boolean] =
+ value(obj, name) match {
case Some(x: Boolean) => Some(x)
case _ => None
}
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/Pure/General/logger.scala Wed Jan 04 14:26:19 2017 +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/General/name_space.ML Wed Jan 04 14:26:08 2017 +0100
+++ b/src/Pure/General/name_space.ML Wed Jan 04 14:26:19 2017 +0100
@@ -213,13 +213,13 @@
(* extern *)
-val names_long_raw = Config.declare_option ("names_long", @{here});
+val names_long_raw = Config.declare_option ("names_long", \<^here>);
val names_long = Config.bool names_long_raw;
-val names_short_raw = Config.declare_option ("names_short", @{here});
+val names_short_raw = Config.declare_option ("names_short", \<^here>);
val names_short = Config.bool names_short_raw;
-val names_unique_raw = Config.declare_option ("names_unique", @{here});
+val names_unique_raw = Config.declare_option ("names_unique", \<^here>);
val names_unique = Config.bool names_unique_raw;
fun extern ctxt space name =
--- a/src/Pure/General/position.ML Wed Jan 04 14:26:08 2017 +0100
+++ b/src/Pure/General/position.ML Wed Jan 04 14:26:19 2017 +0100
@@ -208,7 +208,7 @@
(SOME i, NONE) => (" ", "(line " ^ Value.print_int i ^ ")")
| (SOME i, SOME name) => (" ", "(line " ^ Value.print_int i ^ " of " ^ quote name ^ ")")
| (NONE, SOME name) => (" ", "(file " ^ quote name ^ ")")
- | _ => if is_reported pos then ("", "\<here>") else ("", ""));
+ | _ => if is_reported pos then ("", "\092<^here>") else ("", ""));
in
if null props then ""
else s1 ^ Markup.markup (Markup.properties props Markup.position) s2
--- a/src/Pure/General/position.scala Wed Jan 04 14:26:08 2017 +0100
+++ b/src/Pure/General/position.scala Wed Jan 04 14:26:19 2017 +0100
@@ -64,20 +64,50 @@
}
}
- object Id_Offset0
+ object Item_Id
{
- def unapply(pos: T): Option[(Long, Symbol.Offset)] =
+ def unapply(pos: T): Option[(Long, Symbol.Range)] =
pos match {
- case Id(id) => Some((id, Offset.unapply(pos) getOrElse 0))
+ case Id(id) =>
+ val offset = Offset.unapply(pos) getOrElse 0
+ val end_offset = End_Offset.unapply(pos) getOrElse (offset + 1)
+ Some((id, Text.Range(offset, end_offset)))
case _ => None
}
}
- object Def_Id_Offset0
+ object Item_Def_Id
+ {
+ def unapply(pos: T): Option[(Long, Symbol.Range)] =
+ pos match {
+ case Def_Id(id) =>
+ val offset = Def_Offset.unapply(pos) getOrElse 0
+ val end_offset = Def_End_Offset.unapply(pos) getOrElse (offset + 1)
+ Some((id, Text.Range(offset, end_offset)))
+ case _ => None
+ }
+ }
+
+ object Item_File
{
- def unapply(pos: T): Option[(Long, Symbol.Offset)] =
+ def unapply(pos: T): Option[(String, Int, Symbol.Range)] =
pos match {
- case Def_Id(id) => Some((id, Def_Offset.unapply(pos) getOrElse 0))
+ case Line_File(line, name) =>
+ val offset = Offset.unapply(pos) getOrElse 0
+ val end_offset = End_Offset.unapply(pos) getOrElse (offset + 1)
+ Some((name, line, Text.Range(offset, end_offset)))
+ case _ => None
+ }
+ }
+
+ object Item_Def_File
+ {
+ def unapply(pos: T): Option[(String, Int, Symbol.Range)] =
+ pos match {
+ case Def_Line_File(line, name) =>
+ val offset = Def_Offset.unapply(pos) getOrElse 0
+ val end_offset = Def_End_Offset.unapply(pos) getOrElse (offset + 1)
+ Some((name, line, Text.Range(offset, end_offset)))
case _ => None
}
}
@@ -102,21 +132,20 @@
/* here: user output */
- def here(pos: T): String =
- Markup(Markup.POSITION, pos).markup(
- (Line.unapply(pos), File.unapply(pos)) match {
- case (Some(i), None) => " (line " + i.toString + ")"
- case (Some(i), Some(name)) => " (line " + i.toString + " of " + quote(name) + ")"
- case (None, Some(name)) => " (file " + quote(name) + ")"
- case _ => ""
- })
-
- def here_undelimited(pos: T): String =
- Markup(Markup.POSITION, pos).markup(
- (Line.unapply(pos), File.unapply(pos)) match {
- case (Some(i), None) => "line " + i.toString
- case (Some(i), Some(name)) => "line " + i.toString + " of " + quote(name)
- case (None, Some(name)) => "file " + quote(name)
- case _ => ""
- })
+ def here(props: T, delimited: Boolean = false): String =
+ {
+ val pos = props.filter(p => Markup.POSITION_PROPERTIES(p._1))
+ if (pos.isEmpty) ""
+ else {
+ val s0 =
+ (Line.unapply(pos), File.unapply(pos)) match {
+ case (Some(i), None) => "line " + i.toString
+ case (Some(i), Some(name)) => "line " + i.toString + " of " + quote(name)
+ case (None, Some(name)) => "file " + quote(name)
+ case _ => ""
+ }
+ val s = if (s0 == "") s0 else if (delimited) " (" + s0 + ")" else " " + s0
+ Markup(Markup.POSITION, pos).markup(s)
+ }
+ }
}
--- a/src/Pure/General/source.ML Wed Jan 04 14:26:08 2017 +0100
+++ b/src/Pure/General/source.ML Wed Jan 04 14:26:19 2017 +0100
@@ -14,9 +14,8 @@
val map_filter: ('a -> 'b option) -> ('a, 'c) source -> ('b, ('a, 'c) source) source
val filter: ('a -> bool) -> ('a, 'b) source -> ('a, ('a, 'b) source) source
val of_list: 'a list -> ('a, 'a list) source
+ val of_string: string -> (string, string list) source
val exhausted: ('a, 'b) source -> ('a, 'a list) source
- val of_string: string -> (string, string list) source
- val of_string_limited: int -> string -> (string, substring) source
val source': 'a -> 'b Scan.stopper -> ('a * 'b list -> 'c list * ('a * 'b list)) ->
('b, 'e) source -> ('c, 'a * ('b, 'e) source) source
val source: 'a Scan.stopper -> ('a list -> 'b list * 'a list) ->
@@ -84,20 +83,9 @@
fun of_list xs = make_source [] xs (fn xs => (xs, []));
-fun exhausted src = of_list (exhaust src);
-
-
-(* string source *)
-
val of_string = of_list o raw_explode;
-fun of_string_limited limit str =
- make_source [] (Substring.full str)
- (fn s =>
- let
- val (s1, s2) = Substring.splitAt (s, Int.min (Substring.size s, limit));
- val cs = map String.str (Substring.explode s1);
- in (cs, s2) end);
+fun exhausted src = of_list (exhaust src);
--- a/src/Pure/General/symbol.scala Wed Jan 04 14:26:08 2017 +0100
+++ b/src/Pure/General/symbol.scala Wed Jan 04 14:26:19 2017 +0100
@@ -128,14 +128,15 @@
def explode(text: CharSequence): List[Symbol] = iterator(text).toList
- def advance_line_column(pos: (Int, Int), text: CharSequence): (Int, Int) =
+ def length(text: CharSequence): Int = iterator(text).length
+
+ object Length extends Text.Length
{
- var (line, column) = pos
- for (sym <- iterator(text)) {
- if (is_newline(sym)) { line += 1; column = 1 }
- else column += 1
- }
- (line, column)
+ def apply(text: String): Int = length(text)
+ def offset(text: String, i: Int): Option[Text.Offset] =
+ if (i <= 0 || iterator(text).forall(s => s.length == 1)) Text.Length.offset(text, i)
+ else if (i <= length(text)) Some(iterator(text).take(i).mkString.length)
+ else None
}
--- a/src/Pure/General/url.scala Wed Jan 04 14:26:08 2017 +0100
+++ b/src/Pure/General/url.scala Wed Jan 04 14:26:19 2017 +0100
@@ -7,6 +7,8 @@
package isabelle
+import java.io.{File => JFile}
+import java.net.{URI, URISyntaxException}
import java.net.{URL, MalformedURLException}
import java.util.zip.GZIPInputStream
@@ -45,4 +47,40 @@
def read(name: String): String = read(Url(name), false)
def read_gzip(name: String): String = read(Url(name), true)
+
+
+ /* file URIs */
+
+ def file(uri: String): JFile = new JFile(new URI(uri))
+
+ def is_wellformed_file(uri: String): Boolean =
+ try { file(uri); true }
+ catch { case _: URISyntaxException | _: IllegalArgumentException => false }
+
+ def normalize_file(uri: String): String =
+ if (is_wellformed_file(uri)) {
+ val uri1 = new URI(uri).normalize.toASCIIString
+ if (uri1.startsWith("file://")) uri1
+ else {
+ Library.try_unprefix("file:/", uri1) match {
+ case Some(p) => "file:///" + p
+ case None => uri1
+ }
+ }
+ }
+ else uri
+
+ def platform_file(path: Path): String =
+ {
+ val path1 = path.expand
+ require(path1.is_absolute)
+ platform_file(File.platform_path(path1))
+ }
+
+ def platform_file(name: String): String =
+ if (name.startsWith("file://")) name
+ else {
+ val s = name.replaceAll(" ", "%20")
+ "file://" + (if (Platform.is_windows) s.replace('\\', '/') else s)
+ }
}
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/Pure/General/utf8.scala Wed Jan 04 14:26:19 2017 +0100
@@ -0,0 +1,99 @@
+/* Title: Pure/General/utf8.scala
+ Author: Makarius
+
+Variations on UTF-8.
+*/
+
+package isabelle
+
+
+import java.nio.charset.Charset
+import scala.io.Codec
+
+
+object UTF8
+{
+ /* charset */
+
+ val charset_name: String = "UTF-8"
+ val charset: Charset = Charset.forName(charset_name)
+ def codec(): Codec = Codec(charset)
+
+ def bytes(s: String): Array[Byte] = s.getBytes(charset)
+
+ object Length extends Codepoint.Length
+ {
+ override def codepoint_length(c: Int): Int =
+ if (c < 0x80) 1
+ else if (c < 0x800) 2
+ else if (c < 0x10000) 3
+ else 4
+ }
+
+
+ /* permissive UTF-8 decoding */
+
+ // see also http://en.wikipedia.org/wiki/UTF-8#Description
+ // overlong encodings enable byte-stuffing of low-ASCII
+
+ def decode_permissive(text: CharSequence): String =
+ {
+ val buf = new java.lang.StringBuilder(text.length)
+ var code = -1
+ var rest = 0
+ def flush()
+ {
+ if (code != -1) {
+ if (rest == 0 && Character.isValidCodePoint(code))
+ buf.appendCodePoint(code)
+ else buf.append('\uFFFD')
+ code = -1
+ rest = 0
+ }
+ }
+ def init(x: Int, n: Int)
+ {
+ flush()
+ code = x
+ rest = n
+ }
+ def push(x: Int)
+ {
+ if (rest <= 0) init(x, -1)
+ else {
+ code <<= 6
+ code += x
+ rest -= 1
+ }
+ }
+ for (i <- 0 until text.length) {
+ val c = text.charAt(i)
+ if (c < 128) { flush(); buf.append(c) }
+ else if ((c & 0xC0) == 0x80) push(c & 0x3F)
+ else if ((c & 0xE0) == 0xC0) init(c & 0x1F, 1)
+ else if ((c & 0xF0) == 0xE0) init(c & 0x0F, 2)
+ else if ((c & 0xF8) == 0xF0) init(c & 0x07, 3)
+ }
+ flush()
+ buf.toString
+ }
+
+ private class Decode_Chars(decode: String => String,
+ buffer: Array[Byte], start: Int, end: Int) extends CharSequence
+ {
+ def length: Int = end - start
+ def charAt(i: Int): Char = (buffer(start + i).asInstanceOf[Int] & 0xFF).asInstanceOf[Char]
+ def subSequence(i: Int, j: Int): CharSequence =
+ new Decode_Chars(decode, buffer, start + i, start + j)
+
+ // toString with adhoc decoding: abuse of CharSequence interface
+ override def toString: String = decode(decode_permissive(this))
+ }
+
+ def decode_chars(decode: String => String,
+ buffer: Array[Byte], start: Int, end: Int): CharSequence =
+ {
+ require(0 <= start && start <= end && end <= buffer.length)
+ new Decode_Chars(decode, buffer, start, end)
+ }
+}
--- a/src/Pure/General/word.scala Wed Jan 04 14:26:08 2017 +0100
+++ b/src/Pure/General/word.scala Wed Jan 04 14:26:19 2017 +0100
@@ -12,23 +12,6 @@
object Word
{
- /* codepoints */
-
- def codepoint_iterator(str: String): Iterator[Int] =
- new Iterator[Int] {
- var offset = 0
- def hasNext: Boolean = offset < str.length
- def next: Int =
- {
- val c = str.codePointAt(offset)
- offset += Character.charCount(c)
- c
- }
- }
-
- def codepoint(c: Int): String = new String(Array(c), 0, 1)
-
-
/* directionality */
def bidi_detect(str: String): Boolean =
@@ -51,7 +34,7 @@
}
def perhaps_capitalize(str: String): String =
- if (codepoint_iterator(str).forall(c => Character.isLowerCase(c) || Character.isDigit(c)))
+ if (Codepoint.iterator(str).forall(c => Character.isLowerCase(c) || Character.isDigit(c)))
capitalize(str)
else str
@@ -70,10 +53,10 @@
}
def unapply(str: String): Option[Case] =
if (str.nonEmpty) {
- if (codepoint_iterator(str).forall(Character.isLowerCase(_))) Some(Lowercase)
- else if (codepoint_iterator(str).forall(Character.isUpperCase(_))) Some(Uppercase)
+ if (Codepoint.iterator(str).forall(Character.isLowerCase(_))) Some(Lowercase)
+ else if (Codepoint.iterator(str).forall(Character.isUpperCase(_))) Some(Uppercase)
else {
- val it = codepoint_iterator(str)
+ val it = Codepoint.iterator(str)
if (Character.isUpperCase(it.next) && it.forall(Character.isLowerCase(_)))
Some(Capitalized)
else None
--- a/src/Pure/Isar/attrib.ML Wed Jan 04 14:26:08 2017 +0100
+++ b/src/Pure/Isar/attrib.ML Wed Jan 04 14:26:19 2017 +0100
@@ -216,7 +216,7 @@
(* internal attribute *)
val _ = Theory.setup
- (setup (Binding.make ("attribute", @{here}))
+ (setup (Binding.make ("attribute", \<^here>))
(Scan.lift Args.internal_attribute >> Morphism.form)
"internal attribute");
--- a/src/Pure/Isar/document_structure.scala Wed Jan 04 14:26:08 2017 +0100
+++ b/src/Pure/Isar/document_structure.scala Wed Jan 04 14:26:19 2017 +0100
@@ -190,8 +190,7 @@
toks.filterNot(_.is_space) match {
case List(tok) if tok.is_comment =>
val s = tok.source
- if (Word.codepoint_iterator(s).exists(c =>
- Character.isLetter(c) || Character.isDigit(c)))
+ if (Codepoint.iterator(s).exists(c => Character.isLetter(c) || Character.isDigit(c)))
{
if (s.startsWith("(**** ") && s.endsWith(" ****)")) Some(0)
else if (s.startsWith("(*** ") && s.endsWith(" ***)")) Some(1)
--- a/src/Pure/Isar/local_defs.ML Wed Jan 04 14:26:08 2017 +0100
+++ b/src/Pure/Isar/local_defs.ML Wed Jan 04 14:26:19 2017 +0100
@@ -207,7 +207,7 @@
(* unfold object-level rules *)
-val unfold_abs_def_raw = Config.declare ("unfold_abs_def", @{here}) (K (Config.Bool true));
+val unfold_abs_def_raw = Config.declare ("unfold_abs_def", \<^here>) (K (Config.Bool true));
val unfold_abs_def = Config.bool unfold_abs_def_raw;
local
--- a/src/Pure/Isar/method.ML Wed Jan 04 14:26:08 2017 +0100
+++ b/src/Pure/Isar/method.ML Wed Jan 04 14:26:19 2017 +0100
@@ -393,7 +393,7 @@
val _ =
Theory.setup
- (Global_Theory.add_thms_dynamic (Binding.make ("method_facts", @{here}), get_facts_generic));
+ (Global_Theory.add_thms_dynamic (Binding.make ("method_facts", \<^here>), get_facts_generic));
(* method text *)
@@ -497,7 +497,7 @@
val _ = Seq.pull (method ctxt' src' ctxt' [] (ctxt', Goal.protect 0 Drule.dummy_thm));
in map Token.closure src' end;
-val closure = Config.bool (Config.declare ("Method.closure", @{here}) (K (Config.Bool true)));
+val closure = Config.bool (Config.declare ("Method.closure", \<^here>) (K (Config.Bool true)));
fun method_cmd ctxt =
check_src ctxt #>
@@ -608,7 +608,7 @@
(* sections *)
val old_section_parser =
- Config.bool (Config.declare ("Method.old_section_parser", @{here}) (K (Config.Bool false)));
+ Config.bool (Config.declare ("Method.old_section_parser", \<^here>) (K (Config.Bool false)));
local
--- a/src/Pure/Isar/outer_syntax.ML Wed Jan 04 14:26:08 2017 +0100
+++ b/src/Pure/Isar/outer_syntax.ML Wed Jan 04 14:26:19 2017 +0100
@@ -175,7 +175,7 @@
(* parse commands *)
val bootstrap =
- Config.bool (Config.declare ("Outer_Syntax.bootstrap", @{here}) (K (Config.Bool true)));
+ Config.bool (Config.declare ("Outer_Syntax.bootstrap", \<^here>) (K (Config.Bool true)));
local
@@ -322,7 +322,7 @@
(* 'ML' command -- required for bootstrapping Isar *)
val _ =
- command ("ML", @{here}) "ML text within theory or local theory"
+ command ("ML", \<^here>) "ML text within theory or local theory"
(Parse.ML_source >> (fn source =>
Toplevel.generic_theory
(ML_Context.exec (fn () =>
--- a/src/Pure/Isar/outer_syntax.scala Wed Jan 04 14:26:08 2017 +0100
+++ b/src/Pure/Isar/outer_syntax.scala Wed Jan 04 14:26:19 2017 +0100
@@ -154,8 +154,8 @@
val name = cmd.source
val offset =
(0 /: span.takeWhile(_ != cmd)) {
- case (i, tok) => i + Symbol.iterator(tok.source).length }
- val end_offset = offset + Symbol.iterator(name).length
+ case (i, tok) => i + Symbol.length(tok.source) }
+ val end_offset = offset + Symbol.length(name)
val pos = Position.Range(Text.Range(offset, end_offset) + 1)
Command_Span.Command_Span(name, pos)
}
--- a/src/Pure/Isar/overloading.ML Wed Jan 04 14:26:08 2017 +0100
+++ b/src/Pure/Isar/overloading.ML Wed Jan 04 14:26:19 2017 +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.ML Wed Jan 04 14:26:08 2017 +0100
+++ b/src/Pure/Isar/proof.ML Wed Jan 04 14:26:19 2017 +0100
@@ -522,8 +522,7 @@
fun err_lost () = error ("Lost goal structure:\n" ^ Thm.string_of_thm ctxt goal);
val th =
- (Goal.conclude (if length (flat propss) > 1 then Thm.close_derivation goal else goal)
- handle THM _ => err_lost ())
+ (Goal.conclude (Thm.close_derivation goal) handle THM _ => err_lost ())
|> Drule.flexflex_unique (SOME ctxt)
|> Thm.check_shyps ctxt
|> Thm.check_hyps (Context.Proof ctxt);
--- a/src/Pure/Isar/proof_context.ML Wed Jan 04 14:26:08 2017 +0100
+++ b/src/Pure/Isar/proof_context.ML Wed Jan 04 14:26:19 2017 +0100
@@ -633,7 +633,7 @@
end;
-val show_abbrevs_raw = Config.declare ("show_abbrevs", @{here}) (fn _ => Config.Bool true);
+val show_abbrevs_raw = Config.declare ("show_abbrevs", \<^here>) (fn _ => Config.Bool true);
val show_abbrevs = Config.bool show_abbrevs_raw;
fun contract_abbrevs ctxt t =
@@ -667,7 +667,7 @@
local
val dummies =
- Config.bool (Config.declare ("Proof_Context.dummies", @{here}) (K (Config.Bool false)));
+ Config.bool (Config.declare ("Proof_Context.dummies", \<^here>) (K (Config.Bool false)));
fun check_dummies ctxt t =
if Config.get ctxt dummies then t
@@ -979,7 +979,7 @@
(* retrieve facts *)
val dynamic_facts_dummy =
- Config.bool (Config.declare ("dynamic_facts_dummy_", @{here}) (fn _ => Config.Bool false));
+ Config.bool (Config.declare ("dynamic_facts_dummy_", \<^here>) (fn _ => Config.Bool false));
local
@@ -1049,7 +1049,7 @@
(* inner statement mode *)
val inner_stmt =
- Config.bool (Config.declare ("inner_stmt", @{here}) (K (Config.Bool false)));
+ Config.bool (Config.declare ("inner_stmt", \<^here>) (K (Config.Bool false)));
fun is_stmt ctxt = Config.get ctxt inner_stmt;
val set_stmt = Config.put inner_stmt;
@@ -1571,10 +1571,10 @@
(* core context *)
val debug =
- Config.bool (Config.declare ("Proof_Context.debug", @{here}) (K (Config.Bool false)));
+ Config.bool (Config.declare ("Proof_Context.debug", \<^here>) (K (Config.Bool false)));
val verbose =
- Config.bool (Config.declare ("Proof_Context.verbose", @{here}) (K (Config.Bool false)));
+ Config.bool (Config.declare ("Proof_Context.verbose", \<^here>) (K (Config.Bool false)));
fun pretty_ctxt ctxt =
if not (Config.get ctxt debug) then []
--- a/src/Pure/Isar/proof_display.ML Wed Jan 04 14:26:08 2017 +0100
+++ b/src/Pure/Isar/proof_display.ML Wed Jan 04 14:26:19 2017 +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/Isar/token.ML Wed Jan 04 14:26:08 2017 +0100
+++ b/src/Pure/Isar/token.ML Wed Jan 04 14:26:19 2017 +0100
@@ -275,7 +275,7 @@
| Verbatim => (Markup.verbatim, "")
| Cartouche => (Markup.cartouche, "")
| Comment => (Markup.comment, "")
- | Error msg => (Markup.bad, msg)
+ | Error msg => (Markup.bad (), msg)
| _ => (Markup.empty, "");
fun keyword_reports tok = map (fn markup => ((pos_of tok, markup), ""));
--- a/src/Pure/Isar/token.scala Wed Jan 04 14:26:08 2017 +0100
+++ b/src/Pure/Isar/token.scala Wed Jan 04 14:26:19 2017 +0100
@@ -152,6 +152,8 @@
(toks.toList, ctxt)
}
+ val newline: Token = explode(Keyword.Keywords.empty, "\n").head
+
/* implode */
@@ -205,7 +207,7 @@
def position(): Position.T = position(0)
def position(token: Token): Position.T = position(advance(token).offset)
- override def toString: String = Position.here_undelimited(position())
+ override def toString: String = Position.here(position(), delimited = false)
}
abstract class Reader extends scala.util.parsing.input.Reader[Token]
--- a/src/Pure/ML/ml_antiquotation.ML Wed Jan 04 14:26:08 2017 +0100
+++ b/src/Pure/ML/ml_antiquotation.ML Wed Jan 04 14:26:19 2017 +0100
@@ -34,16 +34,16 @@
(* basic antiquotations *)
val _ = Theory.setup
- (declaration (Binding.make ("here", @{here})) (Scan.succeed ())
+ (declaration (Binding.make ("here", \<^here>)) (Scan.succeed ())
(fn src => fn () =>
ML_Context.value_decl "position" (ML_Syntax.print_position (#2 (Token.name_of_src src)))) #>
- inline (Binding.make ("make_string", @{here})) (Args.context >> K ML_Pretty.make_string_fn) #>
+ inline (Binding.make ("make_string", \<^here>)) (Args.context >> K ML_Pretty.make_string_fn) #>
- value (Binding.make ("binding", @{here}))
+ value (Binding.make ("binding", \<^here>))
(Scan.lift (Parse.position Args.name) >> ML_Syntax.make_binding) #>
- value (Binding.make ("cartouche", @{here}))
+ value (Binding.make ("cartouche", \<^here>))
(Scan.lift Args.cartouche_input >> (fn source =>
"Input.source true " ^ ML_Syntax.print_string (Input.text_of source) ^ " " ^
ML_Syntax.atomic (ML_Syntax.print_range (Input.range_of source)))));
--- a/src/Pure/ML/ml_antiquotations.ML Wed Jan 04 14:26:08 2017 +0100
+++ b/src/Pure/ML/ml_antiquotations.ML Wed Jan 04 14:26:19 2017 +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/ML/ml_compiler.ML Wed Jan 04 14:26:08 2017 +0100
+++ b/src/Pure/ML/ml_compiler.ML Wed Jan 04 14:26:19 2017 +0100
@@ -86,7 +86,7 @@
let
val pos = Exn_Properties.position_of_polyml_location loc;
in
- is_reported pos ?
+ (is_reported pos andalso id <> 0) ?
let
fun markup () =
(Markup.entityN, [(if def then Markup.defN else Markup.refN, Value.print_int id)]);
@@ -109,8 +109,6 @@
| reported loc (PolyML.PTrefId id) = reported_entity_id false (FixedInt.toLarge id) loc
| reported loc (PolyML.PTtype types) = reported_types loc types
| reported loc (PolyML.PTdeclaredAt decl) = reported_entity Markup.ML_defN loc decl
- | reported loc (PolyML.PTopenedAt decl) = reported_entity Markup.ML_openN loc decl
- | reported loc (PolyML.PTstructureAt decl) = reported_entity Markup.ML_structureN loc decl
| reported loc (PolyML.PTcompletions names) = reported_completions loc names
| reported _ _ = I
and reported_tree (loc, props) = fold (reported loc) props;
--- a/src/Pure/ML/ml_compiler0.ML Wed Jan 04 14:26:08 2017 +0100
+++ b/src/Pure/ML/ml_compiler0.ML Wed Jan 04 14:26:19 2017 +0100
@@ -63,7 +63,7 @@
fun ml_input start_line name txt =
let
- fun input line (#"@" :: #"{" :: #"h" :: #"e" :: #"r" :: #"e" :: #"}" :: cs) res =
+ fun input line (#"\\" :: #"<" :: #"^" :: #"h" :: #"e" :: #"r" :: #"e" :: #">" :: cs) res =
let val s = "(Position.line_file_only " ^ Int.toString line ^ " \"" ^ name ^ "\")"
in input line cs (s :: res) end
| input line (#"@" :: #"{" :: #"m" :: #"a" :: #"k" :: #"e" :: #"_" ::
--- a/src/Pure/ML/ml_env.ML Wed Jan 04 14:26:08 2017 +0100
+++ b/src/Pure/ML/ml_env.ML Wed Jan 04 14:26:19 2017 +0100
@@ -98,7 +98,7 @@
(* SML environment for Isabelle/ML *)
val SML_environment =
- Config.bool (Config.declare ("SML_environment", @{here}) (fn _ => Config.Bool false));
+ Config.bool (Config.declare ("SML_environment", \<^here>) (fn _ => Config.Bool false));
fun sml_env SML =
SML orelse
--- a/src/Pure/ML/ml_lex.ML Wed Jan 04 14:26:08 2017 +0100
+++ b/src/Pure/ML/ml_lex.ML Wed Jan 04 14:26:19 2017 +0100
@@ -147,7 +147,7 @@
| Char => (Markup.ML_char, "")
| String => (if SML then Markup.SML_string else Markup.ML_string, "")
| Comment => (if SML then Markup.SML_comment else Markup.ML_comment, "")
- | Error msg => (Markup.bad, msg)
+ | Error msg => (Markup.bad (), msg)
| _ => (Markup.empty, "");
in
--- a/src/Pure/ML/ml_options.ML Wed Jan 04 14:26:08 2017 +0100
+++ b/src/Pure/ML/ml_options.ML Wed Jan 04 14:26:19 2017 +0100
@@ -25,13 +25,13 @@
(* source trace *)
val source_trace_raw =
- Config.declare ("ML_source_trace", @{here}) (fn _ => Config.Bool false);
+ Config.declare ("ML_source_trace", \<^here>) (fn _ => Config.Bool false);
val source_trace = Config.bool source_trace_raw;
(* exception trace *)
-val exception_trace_raw = Config.declare_option ("ML_exception_trace", @{here});
+val exception_trace_raw = Config.declare_option ("ML_exception_trace", \<^here>);
val exception_trace = Config.bool exception_trace_raw;
fun exception_trace_enabled NONE =
@@ -41,7 +41,7 @@
(* exception debugger *)
-val exception_debugger_raw = Config.declare_option ("ML_exception_debugger", @{here});
+val exception_debugger_raw = Config.declare_option ("ML_exception_debugger", \<^here>);
val exception_debugger = Config.bool exception_debugger_raw;
fun exception_debugger_enabled NONE =
@@ -51,7 +51,7 @@
(* debugger *)
-val debugger_raw = Config.declare_option ("ML_debugger", @{here});
+val debugger_raw = Config.declare_option ("ML_debugger", \<^here>);
val debugger = Config.bool debugger_raw;
fun debugger_enabled NONE =
--- a/src/Pure/ML/ml_print_depth.ML Wed Jan 04 14:26:08 2017 +0100
+++ b/src/Pure/ML/ml_print_depth.ML Wed Jan 04 14:26:19 2017 +0100
@@ -18,7 +18,7 @@
val set_print_depth = ML_Print_Depth.set_print_depth;
val print_depth_raw =
- Config.declare ("ML_print_depth", @{here})
+ Config.declare ("ML_print_depth", \<^here>)
(fn _ => Config.Int (ML_Print_Depth.get_print_depth ()));
val print_depth = Config.int print_depth_raw;
--- a/src/Pure/PIDE/command.ML Wed Jan 04 14:26:08 2017 +0100
+++ b/src/Pure/PIDE/command.ML Wed Jan 04 14:26:19 2017 +0100
@@ -116,7 +116,7 @@
Input.source_explode (Token.input_of tok)
|> map_filter (fn (sym, pos) =>
if Symbol.is_malformed sym
- then SOME ((pos, Markup.bad), "Malformed symbolic character") else NONE);
+ then SOME ((pos, Markup.bad ()), "Malformed symbolic character") else NONE);
val is_malformed = Token.is_error tok orelse not (null malformed_symbols);
val reports = Token.reports keywords tok @ Token.completion_report tok @ malformed_symbols;
in (is_malformed, reports) end;
@@ -242,7 +242,7 @@
let
val _ = status tr Markup.failed;
val _ = status tr Markup.finished;
- val _ = if null errs then (report tr Markup.bad; Exn.interrupt ()) else ();
+ val _ = if null errs then (report tr (Markup.bad ()); Exn.interrupt ()) else ();
in {failed = true, command = tr, state = st} end
| SOME st' =>
let
--- a/src/Pure/PIDE/document.scala Wed Jan 04 14:26:08 2017 +0100
+++ b/src/Pure/PIDE/document.scala Wed Jan 04 14:26:19 2017 +0100
@@ -454,6 +454,10 @@
val load_commands: List[Command]
def is_loaded: Boolean
+ def find_command(id: Document_ID.Generic): Option[(Node, Command)]
+ def find_command_position(id: Document_ID.Generic, offset: Symbol.Offset)
+ : Option[Line.Node_Position]
+
def cumulate[A](
range: Text.Range,
info: A,
@@ -546,15 +550,6 @@
def defined_command(id: Document_ID.Command): Boolean = commands.isDefinedAt(id)
- def find_command(version: Version, id: Document_ID.Generic): Option[(Node, Command)] =
- commands.get(id) orElse execs.get(id) match {
- case None => None
- case Some(st) =>
- val command = st.command
- val node = version.nodes(command.node_name)
- if (node.commands.contains(command)) Some((node, command)) else None
- }
-
def the_version(id: Document_ID.Version): Version = versions.getOrElse(id, fail)
def the_static_state(id: Document_ID.Command): Command.State = commands.getOrElse(id, fail)
def the_dynamic_state(id: Document_ID.Exec): Command.State = execs.getOrElse(id, fail)
@@ -793,6 +788,31 @@
val is_loaded: Boolean = node_name.is_theory || load_commands.nonEmpty
+ /* find command */
+
+ def find_command(id: Document_ID.Generic): Option[(Node, Command)] =
+ state.commands.get(id) orElse state.execs.get(id) match {
+ case None => None
+ case Some(st) =>
+ val command = st.command
+ val node = version.nodes(command.node_name)
+ if (node.commands.contains(command)) Some((node, command)) else None
+ }
+
+ def find_command_position(id: Document_ID.Generic, offset: Symbol.Offset)
+ : Option[Line.Node_Position] =
+ for ((node, command) <- find_command(id))
+ yield {
+ val name = command.node_name.node
+ val sources_iterator =
+ node.commands.iterator.takeWhile(_ != command).map(_.source) ++
+ (if (offset == 0) Iterator.empty
+ else Iterator.single(command.source(Text.Range(0, command.chunk.decode(offset)))))
+ val pos = (Line.Position.zero /: sources_iterator)(_.advance(_, Text.Length))
+ Line.Node_Position(name, pos)
+ }
+
+
/* cumulate markup */
def cumulate[A](
@@ -853,4 +873,3 @@
}
}
}
-
--- a/src/Pure/PIDE/editor.scala Wed Jan 04 14:26:08 2017 +0100
+++ b/src/Pure/PIDE/editor.scala Wed Jan 04 14:26:19 2017 +0100
@@ -24,11 +24,11 @@
def remove_overlay(command: Command, fn: String, args: List[String]): Unit
abstract class Hyperlink {
- def external: Boolean
+ def external: Boolean = false
def follow(context: Context): Unit
}
def hyperlink_command(
- focus: Boolean, snapshot: Document.Snapshot, command: Command, offset: Symbol.Offset = 0)
+ focus: Boolean, snapshot: Document.Snapshot, id: Document_ID.Generic, offset: Symbol.Offset = 0)
: Option[Hyperlink]
}
--- a/src/Pure/PIDE/execution.ML Wed Jan 04 14:26:08 2017 +0100
+++ b/src/Pure/PIDE/execution.ML Wed Jan 04 14:26:19 2017 +0100
@@ -116,7 +116,7 @@
Exn.Exn exn =>
(status task [Markup.failed];
status task [Markup.finished];
- Output.report [Markup.markup_only Markup.bad];
+ Output.report [Markup.markup_only (Markup.bad ())];
if exec_id = 0 then ()
else List.app (Future.error_message pos) (Runtime.exn_messages_ids exn))
| Exn.Res _ =>
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/Pure/PIDE/line.scala Wed Jan 04 14:26:19 2017 +0100
@@ -0,0 +1,159 @@
+/* Title: Pure/PIDE/line.scala
+ Author: Makarius
+
+Line-oriented text.
+*/
+
+package isabelle
+
+
+import scala.annotation.tailrec
+
+
+object Line
+{
+ /* position */
+
+ object Position
+ {
+ val zero: Position = Position()
+ }
+
+ sealed case class Position(line: Int = 0, column: Int = 0)
+ {
+ 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
+ }
+
+ def advance(text: String, text_length: Text.Length): Position =
+ if (text.isEmpty) this
+ else {
+ val lines = Library.split_lines(text)
+ val l = line + lines.length - 1
+ val c = (if (l == line) column else 0) + text_length(Library.trim_line(lines.last))
+ Position(l, c)
+ }
+ }
+
+
+ /* range (right-open interval) */
+
+ object Range
+ {
+ def apply(start: Position): Range = Range(start, start)
+ val zero: Range = Range(Position.zero)
+ }
+
+ sealed case class Range(start: Position, stop: Position)
+ {
+ if (start.compare(stop) > 0)
+ error("Bad line range: " + start.print + ".." + stop.print)
+
+ def print: String =
+ if (start == stop) start.print
+ else start.print + ".." + stop.print
+ }
+
+
+ /* positions within document node */
+
+ sealed case class Node_Position(name: String, pos: Position = Position.zero)
+ {
+ def line: Int = pos.line
+ def column: Int = pos.column
+ }
+
+ sealed case class Node_Range(name: String, range: Range = Range.zero)
+ {
+ def start: Position = range.start
+ def stop: Position = range.stop
+ }
+
+
+ /* document with newline as separator (not terminator) */
+
+ object Document
+ {
+ def apply(text: String, text_length: Text.Length): Document =
+ if (text.contains('\r'))
+ Document(Library.split_lines(text).map(s => Line(Library.trim_line(s))), text_length)
+ else
+ Document(Library.split_lines(text).map(s => Line(s)), text_length)
+ }
+
+ sealed case class Document(lines: List[Line], text_length: Text.Length)
+ {
+ def make_text: String = lines.mkString("", "\n", "")
+
+ override def toString: String = make_text
+
+ override def equals(that: Any): Boolean =
+ that match {
+ case other: Document => lines == other.lines
+ case _ => false
+ }
+ override def hashCode(): Int = lines.hashCode
+
+ def position(text_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)
+ case line :: ls =>
+ val n = line.text.length
+ if (ls.isEmpty || i <= n)
+ Position(lines_count).advance(line.text.substring(n - i), text_length)
+ else move(i - (n + 1), lines_count + 1, ls)
+ }
+ }
+ move(text_offset, 0, lines)
+ }
+
+ def range(text_range: Text.Range): Range =
+ Range(position(text_range.start), position(text_range.stop))
+
+ def offset(pos: Position): Option[Text.Offset] =
+ {
+ val l = pos.line
+ val c = pos.column
+ if (0 <= l && l < lines.length) {
+ val line_offset =
+ (0 /: lines.iterator.take(l)) { case (n, line) => n + text_length(line.text) + 1 }
+ text_length.offset(lines(l).text, c).map(line_offset + _)
+ }
+ else None
+ }
+
+ lazy val length: Int =
+ if (lines.isEmpty) 0
+ else ((0 /: lines) { case (n, line) => n + line.text.length + 1 }) - 1
+
+ def full_range: Text.Range = Text.Range(0, length)
+ }
+
+
+ /* 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'))
+
+ 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
+}
--- a/src/Pure/PIDE/markup.ML Wed Jan 04 14:26:08 2017 +0100
+++ b/src/Pure/PIDE/markup.ML Wed Jan 04 14:26:19 2017 +0100
@@ -173,7 +173,7 @@
val protocolN: string
val reportN: string val report: T
val no_reportN: string val no_report: T
- val badN: string val bad: T
+ val badN: string val bad: unit -> T
val intensifyN: string val intensify: T
val browserN: string
val graphviewN: string
@@ -573,7 +573,8 @@
val (reportN, report) = markup_elem "report";
val (no_reportN, no_report) = markup_elem "no_report";
-val (badN, bad) = markup_elem "bad";
+val badN = "bad";
+fun bad () = (badN, serial_properties (serial ()));
val (intensifyN, intensify) = markup_elem "intensify";
--- a/src/Pure/PIDE/protocol.scala Wed Jan 04 14:26:08 2017 +0100
+++ b/src/Pure/PIDE/protocol.scala Wed Jan 04 14:26:19 2017 +0100
@@ -337,8 +337,7 @@
val toks_yxml =
{
import XML.Encode._
- val encode_tok: T[Token] =
- (tok => pair(int, int)((tok.kind.id, Symbol.iterator(tok.source).length)))
+ val encode_tok: T[Token] = (tok => pair(int, int)((tok.kind.id, Symbol.length(tok.source))))
YXML.string_of_body(list(encode_tok)(toks))
}
--- a/src/Pure/PIDE/query_operation.scala Wed Jan 04 14:26:08 2017 +0100
+++ b/src/Pure/PIDE/query_operation.scala Wed Jan 04 14:26:19 2017 +0100
@@ -208,7 +208,7 @@
for {
command <- state.location
snapshot = editor.node_snapshot(command.node_name)
- link <- editor.hyperlink_command(true, snapshot, command)
+ link <- editor.hyperlink_command(true, snapshot, command.id)
} link.follow(editor_context)
}
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/Pure/PIDE/rendering.scala Wed Jan 04 14:26:19 2017 +0100
@@ -0,0 +1,216 @@
+/* Title: Pure/PIDE/rendering.scala
+ Author: Makarius
+
+Isabelle-specific implementation of quasi-abstract rendering and
+markup interpretation.
+*/
+
+package isabelle
+
+
+object Rendering
+{
+ /* message priorities */
+
+ val state_pri = 1
+ val writeln_pri = 2
+ val information_pri = 3
+ val tracing_pri = 4
+ val warning_pri = 5
+ val legacy_pri = 6
+ val error_pri = 7
+
+ val message_pri = Map(
+ Markup.STATE -> state_pri,
+ Markup.STATE_MESSAGE -> state_pri,
+ Markup.WRITELN -> writeln_pri,
+ Markup.WRITELN_MESSAGE -> writeln_pri,
+ Markup.INFORMATION -> information_pri,
+ Markup.INFORMATION_MESSAGE -> information_pri,
+ Markup.TRACING -> tracing_pri,
+ Markup.TRACING_MESSAGE -> tracing_pri,
+ Markup.WARNING -> warning_pri,
+ Markup.WARNING_MESSAGE -> warning_pri,
+ Markup.LEGACY -> legacy_pri,
+ Markup.LEGACY_MESSAGE -> legacy_pri,
+ Markup.ERROR -> error_pri,
+ Markup.ERROR_MESSAGE -> error_pri)
+
+
+ /* markup elements */
+
+ 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)
+
+ val caret_focus_elements = Markup.Elements(Markup.ENTITY)
+}
+
+abstract class Rendering(
+ val snapshot: Document.Snapshot,
+ val options: Options,
+ val resources: Resources)
+{
+ override def toString: String = "Rendering(" + snapshot.toString + ")"
+
+
+ /* tooltips */
+
+ def tooltip_margin: Int
+ def timing_threshold: Double
+
+ def tooltips(range: Text.Range): Option[Text.Info[List[XML.Tree]]] =
+ {
+ 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)), Rendering.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 =>
+ 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.Path(name), _))) =>
+ val file = resources.append_file(snapshot.node_name.master_dir, name)
+ val text =
+ if (name == file) "file " + quote(file)
+ else "path " + quote(name) + "\nfile " + quote(file)
+ Some(add(prev, r, (true, XML.Text(text))))
+
+ 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, Rendering.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, Rendering.pretty_typing("ML:", body))))
+
+ case (prev, Text.Info(r, Protocol.ML_Breakpoint(breakpoint))) =>
+ val text =
+ if (Debugger.breakpoint_state(breakpoint)) "breakpoint (enabled)"
+ else "breakpoint (disabled)"
+ Some(add(prev, r, (true, XML.Text(text))))
+
+ 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, _), _))) =>
+ Rendering.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, all_tips))
+ }
+ }
+
+
+ /* caret focus */
+
+ private def entity_focus(range: Text.Range,
+ check: (Boolean, Long) => Boolean = (is_def: Boolean, i: Long) => is_def): Set[Long] =
+ {
+ val results =
+ snapshot.cumulate[Set[Long]](range, Set.empty, Rendering.caret_focus_elements, _ =>
+ {
+ case (serials, Text.Info(_, XML.Elem(Markup(Markup.ENTITY, props), _))) =>
+ props match {
+ case Markup.Entity.Def(i) if check(true, i) => Some(serials + i)
+ case Markup.Entity.Ref(i) if check(false, i) => Some(serials + i)
+ case _ => None
+ }
+ case _ => None
+ })
+ (Set.empty[Long] /: results){ case (s1, Text.Info(_, s2)) => s1 ++ s2 }
+ }
+
+ def caret_focus(caret_range: Text.Range, visible_range: Text.Range): Set[Long] =
+ {
+ val focus_defs = entity_focus(caret_range)
+ if (focus_defs.nonEmpty) focus_defs
+ else {
+ val visible_defs = entity_focus(visible_range)
+ entity_focus(caret_range, (is_def: Boolean, i: Long) => !is_def && visible_defs.contains(i))
+ }
+ }
+
+ def caret_focus_ranges(caret_range: Text.Range, visible_range: Text.Range): List[Text.Range] =
+ {
+ val focus = caret_focus(caret_range, visible_range)
+ if (focus.nonEmpty) {
+ val results =
+ snapshot.cumulate[Boolean](visible_range, false, Rendering.caret_focus_elements, _ =>
+ {
+ case (_, Text.Info(_, XML.Elem(Markup(Markup.ENTITY, props), _))) =>
+ props match {
+ case Markup.Entity.Def(i) if focus(i) => Some(true)
+ case Markup.Entity.Ref(i) if focus(i) => Some(true)
+ case _ => None
+ }
+ })
+ for (info <- results if info.info) yield info.range
+ }
+ else Nil
+ }
+}
--- a/src/Pure/PIDE/resources.scala Wed Jan 04 14:26:08 2017 +0100
+++ b/src/Pure/PIDE/resources.scala Wed Jan 04 14:26:19 2017 +0100
@@ -23,7 +23,8 @@
class Resources(
val loaded_theories: Set[String],
val known_theories: Map[String, Document.Node.Name],
- val base_syntax: Outer_Syntax)
+ val base_syntax: Outer_Syntax,
+ val log: Logger = No_Logger)
{
/* document node names */
@@ -43,13 +44,33 @@
def append(dir: String, source_path: Path): String =
(Path.explode(dir) + source_path).expand.implode
- def with_thy_reader[A](name: Document.Node.Name, f: Reader[Char] => A): A =
+ def append_file(dir: String, raw_name: String): String =
+ if (Path.is_valid(raw_name)) append(dir, Path.explode(raw_name))
+ else raw_name
+
+
+
+ /* source files of Isabelle/ML bootstrap */
+
+ def source_file(raw_name: String): Option[String] =
{
- val path = Path.explode(name.node)
- if (!path.is_file) error("No such file: " + path.toString)
+ if (Path.is_wellformed(raw_name)) {
+ if (Path.is_valid(raw_name)) {
+ def check(p: Path): Option[Path] = if (p.is_file) Some(p) else None
- val reader = Scan.byte_reader(path.file)
- try { f(reader) } finally { reader.close }
+ val path = Path.explode(raw_name)
+ val path1 =
+ if (path.is_absolute || path.is_current) check(path)
+ else {
+ check(Path.explode("~~/src/Pure") + path) orElse
+ (if (Isabelle_System.getenv("ML_SOURCES") == "") None
+ else check(Path.explode("$ML_SOURCES") + path))
+ }
+ Some(File.platform_path(path1 getOrElse path))
+ }
+ else None
+ }
+ else Some(raw_name)
}
@@ -87,6 +108,15 @@
}
}
+ def with_thy_reader[A](name: Document.Node.Name, f: Reader[Char] => A): A =
+ {
+ val path = Path.explode(name.node)
+ if (!path.is_file) error("No such file: " + path.toString)
+
+ val reader = Scan.byte_reader(path.file)
+ try { f(reader) } finally { reader.close }
+ }
+
def check_thy_reader(qualifier: String, node_name: Document.Node.Name,
reader: Reader[Char], start: Token.Pos): Document.Node.Header =
{
@@ -122,6 +152,16 @@
catch { case ERROR(_) => false }
+ /* special header */
+
+ def special_header(name: Document.Node.Name): Option[Document.Node.Header] =
+ if (Thy_Header.is_ml_root(name.theory))
+ Some(Document.Node.Header(List((import_name("", name, Thy_Header.ML_BOOTSTRAP), Position.none))))
+ else if (Thy_Header.is_bootstrap(name.theory))
+ Some(Document.Node.Header(List((import_name("", name, Thy_Header.PURE), Position.none))))
+ else None
+
+
/* document changes */
def parse_change(
@@ -133,4 +173,3 @@
def commit(change: Session.Change) { }
}
-
--- a/src/Pure/PIDE/text.scala Wed Jan 04 14:26:08 2017 +0100
+++ b/src/Pure/PIDE/text.scala Wed Jan 04 14:26:19 2017 +0100
@@ -25,6 +25,7 @@
{
def apply(start: Offset): Range = Range(start, start)
+ val full: Range = apply(0, Integer.MAX_VALUE / 2)
val offside: Range = apply(-1)
object Ordering extends scala.math.Ordering[Text.Range]
@@ -37,9 +38,9 @@
{
// denotation: {start} Un {i. start < i & i < stop}
if (start > stop)
- error("Bad range: [" + start.toString + ":" + stop.toString + "]")
+ error("Bad range: [" + start.toString + ".." + stop.toString + "]")
- override def toString: String = "[" + start.toString + ":" + stop.toString + "]"
+ override def toString: String = "[" + start.toString + ".." + stop.toString + "]"
def length: Int = stop - start
@@ -79,7 +80,7 @@
{
val empty: Perspective = Perspective(Nil)
- def full: Perspective = Perspective(List(Range(0, Integer.MAX_VALUE / 2)))
+ def full: Perspective = Perspective(List(Range.full))
def apply(ranges: Seq[Range]): Perspective =
{
@@ -180,4 +181,32 @@
(rest, remove(i, count, string))
}
}
+
+
+ /* text length wrt. encoding */
+
+ trait Length
+ {
+ def apply(text: String): Int
+ def offset(text: String, i: Int): Option[Text.Offset]
+ }
+
+ object Length extends Length
+ {
+ def apply(text: String): Int = text.length
+ def offset(text: String, i: Int): Option[Text.Offset] =
+ if (0 <= i && i <= text.length) Some(i) else None
+
+ val encodings: List[(String, Length)] =
+ List(
+ "UTF-16" -> Length,
+ "UTF-8" -> UTF8.Length,
+ "codepoints" -> Codepoint.Length,
+ "symbols" -> Symbol.Length)
+
+ def encoding(name: String): Length =
+ encodings.collectFirst({ case (a, length) if name == a => length }) getOrElse
+ error("Bad text length encoding: " + quote(name) +
+ " (expected " + commas_quote(encodings.map(_._1)) + ")")
+ }
}
--- a/src/Pure/Proof/extraction.ML Wed Jan 04 14:26:08 2017 +0100
+++ b/src/Pure/Proof/extraction.ML Wed Jan 04 14:26:19 2017 +0100
@@ -37,13 +37,13 @@
val add_syntax =
Sign.root_path
#> Sign.add_types_global
- [(Binding.make ("Type", @{here}), 0, NoSyn),
- (Binding.make ("Null", @{here}), 0, NoSyn)]
+ [(Binding.make ("Type", \<^here>), 0, NoSyn),
+ (Binding.make ("Null", \<^here>), 0, NoSyn)]
#> Sign.add_consts
- [(Binding.make ("typeof", @{here}), typ "'b => Type", NoSyn),
- (Binding.make ("Type", @{here}), typ "'a itself => Type", NoSyn),
- (Binding.make ("Null", @{here}), typ "Null", NoSyn),
- (Binding.make ("realizes", @{here}), typ "'a => 'b => 'b", NoSyn)];
+ [(Binding.make ("typeof", \<^here>), typ "'b => Type", NoSyn),
+ (Binding.make ("Type", \<^here>), typ "'a itself => Type", NoSyn),
+ (Binding.make ("Null", \<^here>), typ "Null", NoSyn),
+ (Binding.make ("realizes", \<^here>), typ "'a => 'b => 'b", NoSyn)];
val nullT = Type ("Null", []);
val nullt = Const ("Null", nullT);
--- a/src/Pure/Proof/proof_syntax.ML Wed Jan 04 14:26:08 2017 +0100
+++ b/src/Pure/Proof/proof_syntax.ML Wed Jan 04 14:26:19 2017 +0100
@@ -48,19 +48,19 @@
|> Sign.root_path
|> Sign.set_defsort []
|> Sign.add_types_global
- [(Binding.make ("proof", @{here}), 0, NoSyn)]
+ [(Binding.make ("proof", \<^here>), 0, NoSyn)]
|> fold (snd oo Sign.declare_const_global)
- [((Binding.make ("Appt", @{here}), [proofT, aT] ---> proofT), mixfix ("(1_ \<cdot>/ _)", [4, 5], 4)),
- ((Binding.make ("AppP", @{here}), [proofT, proofT] ---> proofT), mixfix ("(1_ \<bullet>/ _)", [4, 5], 4)),
- ((Binding.make ("Abst", @{here}), (aT --> proofT) --> proofT), NoSyn),
- ((Binding.make ("AbsP", @{here}), [propT, proofT --> proofT] ---> proofT), NoSyn),
- ((Binding.make ("Hyp", @{here}), propT --> proofT), NoSyn),
- ((Binding.make ("Oracle", @{here}), propT --> proofT), NoSyn),
- ((Binding.make ("OfClass", @{here}), (Term.a_itselfT --> propT) --> proofT), NoSyn),
- ((Binding.make ("MinProof", @{here}), proofT), Mixfix.mixfix "?")]
+ [((Binding.make ("Appt", \<^here>), [proofT, aT] ---> proofT), mixfix ("(1_ \<cdot>/ _)", [4, 5], 4)),
+ ((Binding.make ("AppP", \<^here>), [proofT, proofT] ---> proofT), mixfix ("(1_ \<bullet>/ _)", [4, 5], 4)),
+ ((Binding.make ("Abst", \<^here>), (aT --> proofT) --> proofT), NoSyn),
+ ((Binding.make ("AbsP", \<^here>), [propT, proofT --> proofT] ---> proofT), NoSyn),
+ ((Binding.make ("Hyp", \<^here>), propT --> proofT), NoSyn),
+ ((Binding.make ("Oracle", \<^here>), propT --> proofT), NoSyn),
+ ((Binding.make ("OfClass", \<^here>), (Term.a_itselfT --> propT) --> proofT), NoSyn),
+ ((Binding.make ("MinProof", \<^here>), proofT), Mixfix.mixfix "?")]
|> Sign.add_nonterminals_global
- [Binding.make ("param", @{here}),
- Binding.make ("params", @{here})]
+ [Binding.make ("param", \<^here>),
+ Binding.make ("params", \<^here>)]
|> Sign.add_syntax Syntax.mode_default
[("_Lam", [paramsT, proofT] ---> proofT, mixfix ("(1\<^bold>\<lambda>_./ _)", [0, 3], 3)),
("_Lam0", [paramT, paramsT] ---> paramsT, mixfix ("_/ _", [1, 0], 0)),
--- a/src/Pure/Proof/reconstruct.ML Wed Jan 04 14:26:08 2017 +0100
+++ b/src/Pure/Proof/reconstruct.ML Wed Jan 04 14:26:19 2017 +0100
@@ -19,7 +19,7 @@
struct
val quiet_mode =
- Config.bool (Config.declare ("Reconstruct.quiet_mode", @{here}) (K (Config.Bool true)));
+ Config.bool (Config.declare ("Reconstruct.quiet_mode", \<^here>) (K (Config.Bool true)));
fun message ctxt msg =
if Config.get ctxt quiet_mode then () else writeln (msg ());
--- a/src/Pure/Pure.thy Wed Jan 04 14:26:08 2017 +0100
+++ b/src/Pure/Pure.thy Wed Jan 04 14:26:19 2017 +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)));
@@ -991,7 +991,7 @@
local
fun report_back () =
- Output.report [Markup.markup Markup.bad "Explicit backtracking"];
+ Output.report [Markup.markup (Markup.bad ()) "Explicit backtracking"];
val _ =
Outer_Syntax.command @{command_keyword back} "explicit backtracking of proof command"
@@ -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/Syntax/ast.ML Wed Jan 04 14:26:08 2017 +0100
+++ b/src/Pure/Syntax/ast.ML Wed Jan 04 14:26:19 2017 +0100
@@ -166,10 +166,10 @@
(* normalize *)
-val trace_raw = Config.declare ("syntax_ast_trace", @{here}) (fn _ => Config.Bool false);
+val trace_raw = Config.declare ("syntax_ast_trace", \<^here>) (fn _ => Config.Bool false);
val trace = Config.bool trace_raw;
-val stats_raw = Config.declare ("syntax_ast_stats", @{here}) (fn _ => Config.Bool false);
+val stats_raw = Config.declare ("syntax_ast_stats", \<^here>) (fn _ => Config.Bool false);
val stats = Config.bool stats_raw;
fun message head body =
--- a/src/Pure/Syntax/parser.ML Wed Jan 04 14:26:08 2017 +0100
+++ b/src/Pure/Syntax/parser.ML Wed Jan 04 14:26:19 2017 +0100
@@ -627,7 +627,7 @@
(*trigger value for warnings*)
val branching_level =
- Config.int (Config.declare ("syntax_branching_level", @{here}) (fn _ => Config.Int 600));
+ Config.int (Config.declare ("syntax_branching_level", \<^here>) (fn _ => Config.Int 600));
(*get all productions of a NT and NTs chained to it which can
be started by specified token*)
--- a/src/Pure/Syntax/printer.ML Wed Jan 04 14:26:08 2017 +0100
+++ b/src/Pure/Syntax/printer.ML Wed Jan 04 14:26:19 2017 +0100
@@ -48,29 +48,29 @@
(** options for printing **)
-val show_brackets_raw = Config.declare_option ("show_brackets", @{here});
+val show_brackets_raw = Config.declare_option ("show_brackets", \<^here>);
val show_brackets = Config.bool show_brackets_raw;
-val show_types_raw = Config.declare_option ("show_types", @{here});
+val show_types_raw = Config.declare_option ("show_types", \<^here>);
val show_types = Config.bool show_types_raw;
-val show_sorts_raw = Config.declare_option ("show_sorts", @{here});
+val show_sorts_raw = Config.declare_option ("show_sorts", \<^here>);
val show_sorts = Config.bool show_sorts_raw;
val show_markup_default = Unsynchronized.ref false;
val show_markup_raw =
- Config.declare ("show_markup", @{here}) (fn _ => Config.Bool (! show_markup_default));
+ Config.declare ("show_markup", \<^here>) (fn _ => Config.Bool (! show_markup_default));
val show_markup = Config.bool show_markup_raw;
val show_structs_raw =
- Config.declare ("show_structs", @{here}) (fn _ => Config.Bool false);
+ Config.declare ("show_structs", \<^here>) (fn _ => Config.Bool false);
val show_structs = Config.bool show_structs_raw;
-val show_question_marks_raw = Config.declare_option ("show_question_marks", @{here});
+val show_question_marks_raw = Config.declare_option ("show_question_marks", \<^here>);
val show_question_marks = Config.bool show_question_marks_raw;
val show_type_emphasis =
- Config.bool (Config.declare ("show_type_emphasis", @{here}) (fn _ => Config.Bool true));
+ Config.bool (Config.declare ("show_type_emphasis", \<^here>) (fn _ => Config.Bool true));
fun type_emphasis ctxt T =
T <> dummyT andalso
@@ -169,7 +169,7 @@
| is_chain _ = false;
val pretty_priority =
- Config.int (Config.declare ("Syntax.pretty_priority", @{here}) (K (Config.Int 0)));
+ Config.int (Config.declare ("Syntax.pretty_priority", \<^here>) (K (Config.Int 0)));
fun pretty type_mode curried ctxt tabs trf markup_trans markup_extern ast0 =
let
--- a/src/Pure/Syntax/syntax.ML Wed Jan 04 14:26:08 2017 +0100
+++ b/src/Pure/Syntax/syntax.ML Wed Jan 04 14:26:19 2017 +0100
@@ -160,14 +160,14 @@
(* configuration options *)
-val root = Config.string (Config.declare ("syntax_root", @{here}) (K (Config.String "any")));
+val root = Config.string (Config.declare ("syntax_root", \<^here>) (K (Config.String "any")));
val ambiguity_warning_raw =
- Config.declare ("syntax_ambiguity_warning", @{here}) (fn _ => Config.Bool true);
+ Config.declare ("syntax_ambiguity_warning", \<^here>) (fn _ => Config.Bool true);
val ambiguity_warning = Config.bool ambiguity_warning_raw;
val ambiguity_limit_raw =
- Config.declare ("syntax_ambiguity_limit", @{here}) (fn _ => Config.Int 10);
+ Config.declare ("syntax_ambiguity_limit", \<^here>) (fn _ => Config.Int 10);
val ambiguity_limit = Config.int ambiguity_limit_raw;
@@ -330,7 +330,7 @@
(* global pretty printing *)
val pretty_global =
- Config.bool (Config.declare ("Syntax.pretty_global", @{here}) (K (Config.Bool false)));
+ Config.bool (Config.declare ("Syntax.pretty_global", \<^here>) (K (Config.Bool false)));
fun is_pretty_global ctxt = Config.get ctxt pretty_global;
val set_pretty_global = Config.put pretty_global;
val init_pretty_global = set_pretty_global true o Proof_Context.init_global;
--- a/src/Pure/Syntax/syntax_ext.ML Wed Jan 04 14:26:08 2017 +0100
+++ b/src/Pure/Syntax/syntax_ext.ML Wed Jan 04 14:26:19 2017 +0100
@@ -217,7 +217,7 @@
(case xsym of
Delim s =>
if Position.is_reported pos andalso exists Symbol.is_utf8 (Symbol.explode s) then
- [((pos, Markup.bad),
+ [((pos, Markup.bad ()),
"Mixfix delimiter contains raw Unicode -- this is non-portable and unreliable")]
else []
| _ => []);
--- a/src/Pure/Syntax/syntax_phases.ML Wed Jan 04 14:26:08 2017 +0100
+++ b/src/Pure/Syntax/syntax_phases.ML Wed Jan 04 14:26:19 2017 +0100
@@ -375,7 +375,7 @@
fun parse_failed ctxt pos msg kind =
cat_error msg ("Failed to parse " ^ kind ^
- Markup.markup_report (Context_Position.reported_text ctxt pos Markup.bad ""));
+ Markup.markup_report (Context_Position.reported_text ctxt pos (Markup.bad ()) ""));
fun parse_sort ctxt =
Syntax.parse_input ctxt Term_XML.Decode.sort Markup.language_sort
@@ -746,7 +746,7 @@
| token_trans "_tvar" x = SOME (Pretty.mark_str (Markup.tvar, x))
| token_trans "_free" x = SOME (Pretty.marks_str (free_or_skolem ctxt x))
| token_trans "_bound" x = SOME (Pretty.mark_str (Markup.bound, x))
- | token_trans "_loose" x = SOME (Pretty.mark_str (Markup.bad, x))
+ | token_trans "_loose" x = SOME (Pretty.mark_str (Markup.bad (), x))
| token_trans "_var" x = SOME (Pretty.mark_str (var_or_skolem x))
| token_trans "_numeral" x = SOME (Pretty.mark_str (Markup.numeral, x))
| token_trans "_inner_string" x = SOME (Pretty.mark_str (Markup.inner_string, x))
--- a/src/Pure/Syntax/syntax_trans.ML Wed Jan 04 14:26:08 2017 +0100
+++ b/src/Pure/Syntax/syntax_trans.ML Wed Jan 04 14:26:19 2017 +0100
@@ -297,7 +297,7 @@
| t' => Abs (a, T, t'))
| eta_abs t = t;
-val eta_contract_raw = Config.declare_option ("eta_contract", @{here});
+val eta_contract_raw = Config.declare_option ("eta_contract", \<^here>);
val eta_contract = Config.bool eta_contract_raw;
fun eta_contr ctxt tm =
--- a/src/Pure/System/isabelle_system.scala Wed Jan 04 14:26:08 2017 +0100
+++ b/src/Pure/System/isabelle_system.scala Wed Jan 04 14:26:19 2017 +0100
@@ -161,21 +161,6 @@
/** file-system operations **/
- /* source files of Isabelle/ML bootstrap */
-
- def source_file(path: Path): Option[Path] =
- {
- def check(p: Path): Option[Path] = if (p.is_file) Some(p) else None
-
- if (path.is_absolute || path.is_current) check(path)
- else {
- check(Path.explode("~~/src/Pure") + path) orElse
- (if (getenv("ML_SOURCES") == "") None
- else check(Path.explode("$ML_SOURCES") + path))
- }
- }
-
-
/* directories */
def mkdirs(path: Path): Unit =
--- a/src/Pure/System/isabelle_tool.scala Wed Jan 04 14:26:08 2017 +0100
+++ b/src/Pure/System/isabelle_tool.scala Wed Jan 04 14:26:19 2017 +0100
@@ -113,7 +113,9 @@
Update_Cartouches.isabelle_tool,
Update_Header.isabelle_tool,
Update_Then.isabelle_tool,
- Update_Theorems.isabelle_tool)
+ Update_Theorems.isabelle_tool,
+ isabelle.vscode.Grammar.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/System/utf8.scala Wed Jan 04 14:26:08 2017 +0100
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,91 +0,0 @@
-/* Title: Pure/System/utf8.scala
- Author: Makarius
-
-Variations on UTF-8.
-*/
-
-package isabelle
-
-
-import java.nio.charset.Charset
-import scala.io.Codec
-
-
-object UTF8
-{
- /* charset */
-
- val charset_name: String = "UTF-8"
- val charset: Charset = Charset.forName(charset_name)
- def codec(): Codec = Codec(charset)
-
- def bytes(s: String): Array[Byte] = s.getBytes(charset)
-
-
- /* permissive UTF-8 decoding */
-
- // see also http://en.wikipedia.org/wiki/UTF-8#Description
- // overlong encodings enable byte-stuffing of low-ASCII
-
- def decode_permissive(text: CharSequence): String =
- {
- val buf = new java.lang.StringBuilder(text.length)
- var code = -1
- var rest = 0
- def flush()
- {
- if (code != -1) {
- if (rest == 0 && Character.isValidCodePoint(code))
- buf.appendCodePoint(code)
- else buf.append('\uFFFD')
- code = -1
- rest = 0
- }
- }
- def init(x: Int, n: Int)
- {
- flush()
- code = x
- rest = n
- }
- def push(x: Int)
- {
- if (rest <= 0) init(x, -1)
- else {
- code <<= 6
- code += x
- rest -= 1
- }
- }
- for (i <- 0 until text.length) {
- val c = text.charAt(i)
- if (c < 128) { flush(); buf.append(c) }
- else if ((c & 0xC0) == 0x80) push(c & 0x3F)
- else if ((c & 0xE0) == 0xC0) init(c & 0x1F, 1)
- else if ((c & 0xF0) == 0xE0) init(c & 0x0F, 2)
- else if ((c & 0xF8) == 0xF0) init(c & 0x07, 3)
- }
- flush()
- buf.toString
- }
-
- private class Decode_Chars(decode: String => String,
- buffer: Array[Byte], start: Int, end: Int) extends CharSequence
- {
- def length: Int = end - start
- def charAt(i: Int): Char = (buffer(start + i).asInstanceOf[Int] & 0xFF).asInstanceOf[Char]
- def subSequence(i: Int, j: Int): CharSequence =
- new Decode_Chars(decode, buffer, start + i, start + j)
-
- // toString with adhoc decoding: abuse of CharSequence interface
- override def toString: String = decode(decode_permissive(this))
- }
-
- def decode_chars(decode: String => String,
- buffer: Array[Byte], start: Int, end: Int): CharSequence =
- {
- require(0 <= start && start <= end && end <= buffer.length)
- new Decode_Chars(decode, buffer, start, end)
- }
-}
-
--- a/src/Pure/Thy/thy_header.ML Wed Jan 04 14:26:08 2017 +0100
+++ b/src/Pure/Thy/thy_header.ML Wed Jan 04 14:26:19 2017 +0100
@@ -64,28 +64,28 @@
val bootstrap_keywords =
Keyword.empty_keywords
|> Keyword.add_keywords
- [(("%", @{here}), Keyword.no_spec),
- (("(", @{here}), Keyword.no_spec),
- ((")", @{here}), Keyword.no_spec),
- ((",", @{here}), Keyword.no_spec),
- (("::", @{here}), Keyword.no_spec),
- (("=", @{here}), Keyword.no_spec),
- (("and", @{here}), Keyword.no_spec),
- ((beginN, @{here}), Keyword.quasi_command_spec),
- ((importsN, @{here}), Keyword.quasi_command_spec),
- ((keywordsN, @{here}), Keyword.quasi_command_spec),
- ((abbrevsN, @{here}), Keyword.quasi_command_spec),
- ((chapterN, @{here}), ((Keyword.document_heading, []), [])),
- ((sectionN, @{here}), ((Keyword.document_heading, []), [])),
- ((subsectionN, @{here}), ((Keyword.document_heading, []), [])),
- ((subsubsectionN, @{here}), ((Keyword.document_heading, []), [])),
- ((paragraphN, @{here}), ((Keyword.document_heading, []), [])),
- ((subparagraphN, @{here}), ((Keyword.document_heading, []), [])),
- ((textN, @{here}), ((Keyword.document_body, []), [])),
- ((txtN, @{here}), ((Keyword.document_body, []), [])),
- ((text_rawN, @{here}), ((Keyword.document_raw, []), [])),
- ((theoryN, @{here}), ((Keyword.thy_begin, []), ["theory"])),
- (("ML", @{here}), ((Keyword.thy_decl, []), ["ML"]))];
+ [(("%", \<^here>), Keyword.no_spec),
+ (("(", \<^here>), Keyword.no_spec),
+ ((")", \<^here>), Keyword.no_spec),
+ ((",", \<^here>), Keyword.no_spec),
+ (("::", \<^here>), Keyword.no_spec),
+ (("=", \<^here>), Keyword.no_spec),
+ (("and", \<^here>), Keyword.no_spec),
+ ((beginN, \<^here>), Keyword.quasi_command_spec),
+ ((importsN, \<^here>), Keyword.quasi_command_spec),
+ ((keywordsN, \<^here>), Keyword.quasi_command_spec),
+ ((abbrevsN, \<^here>), Keyword.quasi_command_spec),
+ ((chapterN, \<^here>), ((Keyword.document_heading, []), [])),
+ ((sectionN, \<^here>), ((Keyword.document_heading, []), [])),
+ ((subsectionN, \<^here>), ((Keyword.document_heading, []), [])),
+ ((subsubsectionN, \<^here>), ((Keyword.document_heading, []), [])),
+ ((paragraphN, \<^here>), ((Keyword.document_heading, []), [])),
+ ((subparagraphN, \<^here>), ((Keyword.document_heading, []), [])),
+ ((textN, \<^here>), ((Keyword.document_body, []), [])),
+ ((txtN, \<^here>), ((Keyword.document_body, []), [])),
+ ((text_rawN, \<^here>), ((Keyword.document_raw, []), [])),
+ ((theoryN, \<^here>), ((Keyword.thy_begin, []), ["theory"])),
+ (("ML", \<^here>), ((Keyword.thy_decl, []), ["ML"]))];
(* theory data *)
--- a/src/Pure/Thy/thy_header.scala Wed Jan 04 14:26:08 2017 +0100
+++ b/src/Pure/Thy/thy_header.scala Wed Jan 04 14:26:19 2017 +0100
@@ -69,7 +69,7 @@
Outer_Syntax.init().add_keywords(bootstrap_header)
- /* file name */
+ /* file name vs. theory name */
val PURE = "Pure"
val ML_BOOTSTRAP = "ML_Bootstrap"
@@ -77,9 +77,13 @@
val ml_roots = List("ROOT0.ML" -> "ML_Root0", "ROOT.ML" -> ML_ROOT)
val bootstrap_thys = List(PURE, ML_BOOTSTRAP).map(a => a -> ("Bootstrap_" + a))
+ private val Dir_Name = new Regex("""(.*?)[^/\\:]+""")
private val Base_Name = new Regex(""".*?([^/\\:]+)""")
private val Thy_Name = new Regex(""".*?([^/\\:]+)\.thy""")
+ def dir_name(s: String): String =
+ s match { case Dir_Name(name) => name case _ => "" }
+
def base_name(s: String): String =
s match { case Base_Name(name) => name case _ => error("Malformed import: " + quote(s)) }
@@ -94,6 +98,12 @@
case _ => None
}
+ def is_ml_root(theory: String): Boolean =
+ ml_roots.exists({ case (_, b) => b == theory })
+
+ def is_bootstrap(theory: String): Boolean =
+ bootstrap_thys.exists({ case (_, b) => b == theory })
+
/* header */
@@ -172,6 +182,35 @@
def read(source: CharSequence, start: Token.Pos): Thy_Header =
read(new CharSequenceReader(source), start)
+
+
+ /* line-oriented text */
+
+ def header_text(doc: Line.Document): String =
+ {
+ val keywords = bootstrap_syntax.keywords
+ val toks = new mutable.ListBuffer[Token]
+ val iterator =
+ (for {
+ (toks, _) <-
+ doc.lines.iterator.scanLeft((List.empty[Token], Scan.Finished: Scan.Line_Context))(
+ {
+ case ((_, ctxt), line) => Token.explode_line(keywords, line.text, ctxt)
+ })
+ tok <- toks.iterator ++ Iterator.single(Token.newline)
+ } yield tok).dropWhile(tok => !tok.is_command(Thy_Header.THEORY))
+
+ @tailrec def until_begin
+ {
+ if (iterator.hasNext) {
+ val tok = iterator.next
+ toks += tok
+ if (!tok.is_begin) until_begin
+ }
+ }
+ until_begin
+ Token.implode(toks.toList)
+ }
}
--- a/src/Pure/Thy/thy_info.ML Wed Jan 04 14:26:08 2017 +0100
+++ b/src/Pure/Thy/thy_info.ML Wed Jan 04 14:26:19 2017 +0100
@@ -159,7 +159,7 @@
(*toplevel proofs and diags*)
val _ = Future.join_tasks (maps Future.group_snapshot (Execution.peek exec_id));
(*fully nested proofs*)
- val res = Exn.capture Thm.join_theory_proofs theory;
+ val res = Exn.capture Thm.consolidate_theory theory;
in res :: map Exn.Exn (maps Task_Queue.group_status (Execution.peek exec_id)) end;
datatype task =
--- a/src/Pure/Thy/thy_output.ML Wed Jan 04 14:26:08 2017 +0100
+++ b/src/Pure/Thy/thy_output.ML Wed Jan 04 14:26:19 2017 +0100
@@ -44,13 +44,13 @@
(** options **)
-val display = Attrib.setup_option_bool ("thy_output_display", @{here});
-val break = Attrib.setup_option_bool ("thy_output_break", @{here});
-val quotes = Attrib.setup_option_bool ("thy_output_quotes", @{here});
-val margin = Attrib.setup_option_int ("thy_output_margin", @{here});
-val indent = Attrib.setup_option_int ("thy_output_indent", @{here});
-val source = Attrib.setup_option_bool ("thy_output_source", @{here});
-val modes = Attrib.setup_option_string ("thy_output_modes", @{here});
+val display = Attrib.setup_option_bool ("thy_output_display", \<^here>);
+val break = Attrib.setup_option_bool ("thy_output_break", \<^here>);
+val quotes = Attrib.setup_option_bool ("thy_output_quotes", \<^here>);
+val margin = Attrib.setup_option_int ("thy_output_margin", \<^here>);
+val indent = Attrib.setup_option_int ("thy_output_indent", \<^here>);
+val source = Attrib.setup_option_bool ("thy_output_source", \<^here>);
+val modes = Attrib.setup_option_string ("thy_output_modes", \<^here>);
structure Wrappers = Proof_Data
--- a/src/Pure/Tools/build.ML Wed Jan 04 14:26:08 2017 +0100
+++ b/src/Pure/Tools/build.ML Wed Jan 04 14:26:19 2017 +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 Wed Jan 04 14:26:08 2017 +0100
+++ b/src/Pure/Tools/build.scala Wed Jan 04 14:26:19 2017 +0100
@@ -95,6 +95,12 @@
/* source dependencies and static content */
+ object Session_Content
+ {
+ def empty: Session_Content =
+ Session_Content(Set.empty, Map.empty, Nil, Outer_Syntax.empty, Nil, Graph_Display.empty_graph)
+ }
+
sealed case class Session_Content(
loaded_theories: Set[String],
known_theories: Map[String, Document.Node.Name],
@@ -229,8 +235,8 @@
session_dependencies(options, inlined_files, dirs, List(session))(session)
}
- def outer_syntax(options: Options, session: String): Outer_Syntax =
- session_content(options, false, Nil, session).syntax
+ def outer_syntax(options: Options, dirs: List[Path], session: String): Outer_Syntax =
+ session_content(options, false, dirs, session).syntax
/* jobs */
--- a/src/Pure/Tools/find_consts.ML Wed Jan 04 14:26:08 2017 +0100
+++ b/src/Pure/Tools/find_consts.ML Wed Jan 04 14:26:19 2017 +0100
@@ -141,7 +141,7 @@
Parse.typ >> Loose;
val query_keywords =
- Keyword.add_keywords [((":", @{here}), Keyword.no_spec)] Keyword.empty_keywords;
+ Keyword.add_keywords [((":", \<^here>), Keyword.no_spec)] Keyword.empty_keywords;
in
--- a/src/Pure/Tools/find_theorems.ML Wed Jan 04 14:26:08 2017 +0100
+++ b/src/Pure/Tools/find_theorems.ML Wed Jan 04 14:26:19 2017 +0100
@@ -522,7 +522,7 @@
Parse.term >> Pattern;
val query_keywords =
- Keyword.add_keywords [((":", @{here}), Keyword.no_spec)] Keyword.empty_keywords;
+ Keyword.add_keywords [((":", \<^here>), Keyword.no_spec)] Keyword.empty_keywords;
in
--- a/src/Pure/Tools/ml_console.scala Wed Jan 04 14:26:08 2017 +0100
+++ b/src/Pure/Tools/ml_console.scala Wed Jan 04 14:26:19 2017 +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/Tools/ml_process.scala Wed Jan 04 14:26:08 2017 +0100
+++ b/src/Pure/Tools/ml_process.scala Wed Jan 04 14:26:19 2017 +0100
@@ -48,7 +48,18 @@
fun subparagraph (_: string) = ();
val ML_file = PolyML.use;
""",
- if (Isabelle_System.getenv("ML_SYSTEM") == "polyml-5.6") "structure FixedInt = IntInf"
+ if (Isabelle_System.getenv("ML_SYSTEM") == "polyml-5.6")
+ """
+ structure FixedInt = IntInf;
+ structure RunCall =
+ struct
+ open RunCall
+ val loadWord: word * word -> word =
+ run_call2 RuntimeCalls.POLY_SYS_load_word;
+ val storeWord: word * word * word -> unit =
+ run_call3 RuntimeCalls.POLY_SYS_assign_word;
+ end;
+ """
else "",
if (Platform.is_windows)
"fun exit 0 = OS.Process.exit OS.Process.success" +
--- a/src/Pure/Tools/print_operation.ML Wed Jan 04 14:26:08 2017 +0100
+++ b/src/Pure/Tools/print_operation.ML Wed Jan 04 14:26:19 2017 +0100
@@ -49,7 +49,7 @@
(fn {state, args, writeln_result, ...} =>
let
val _ = Toplevel.context_of state handle Toplevel.UNDEF => error "Unknown context";
- fun err s = Pretty.mark_str (Markup.bad, s);
+ fun err s = Pretty.mark_str (Markup.bad (), s);
fun print name =
(case AList.lookup (op =) (Synchronized.value print_operations) name of
SOME (_, pr) => (pr state handle Toplevel.UNDEF => [err "Unknown context"])
--- a/src/Pure/Tools/thm_deps.ML Wed Jan 04 14:26:08 2017 +0100
+++ b/src/Pure/Tools/thm_deps.ML Wed Jan 04 14:26:19 2017 +0100
@@ -20,8 +20,8 @@
fun make_node name directory =
Graph_Display.session_node
{name = Long_Name.base_name name, directory = directory, unfold = false, path = ""};
- fun add_dep ("", _, _) = I
- | add_dep (name, _, PBody {thms = thms', ...}) =
+ fun add_dep {name = "", ...} = I
+ | add_dep {name = name, body = PBody {thms = thms', ...}, ...} =
let
val prefix = #1 (split_last (Long_Name.explode name));
val session =
@@ -35,7 +35,7 @@
| NONE => [])
| _ => ["global"]);
val node = make_node name (space_implode "/" (session @ prefix));
- val deps = filter_out (fn s => s = "") (map (#1 o #2) thms');
+ val deps = filter_out (fn s => s = "") (map (Proofterm.thm_node_name o #2) thms');
in Symtab.update (name, (node, deps)) end;
val entries0 = Proofterm.fold_body_thms add_dep (Thm.proof_bodies_of thms) Symtab.empty;
val entries1 =
@@ -73,7 +73,7 @@
val used =
Proofterm.fold_body_thms
- (fn (a, _, _) => a <> "" ? Symtab.update (a, ()))
+ (fn {name = a, ...} => a <> "" ? Symtab.update (a, ()))
(map Proofterm.strip_thm (Thm.proof_bodies_of (map (#1 o #2) new_thms)))
Symtab.empty;
--- a/src/Pure/build-jars Wed Jan 04 14:26:08 2017 +0100
+++ b/src/Pure/build-jars Wed Jan 04 14:26:19 2017 +0100
@@ -39,16 +39,19 @@
GUI/wrap_panel.scala
General/antiquote.scala
General/bytes.scala
+ General/codepoint.scala
General/completion.scala
General/date.scala
General/exn.scala
General/file.scala
+ General/file_watcher.scala
General/graph.scala
General/graph_display.scala
General/graphics_file.scala
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
@@ -67,6 +70,7 @@
General/timing.scala
General/untyped.scala
General/url.scala
+ General/utf8.scala
General/value.scala
General/word.scala
General/xz.scala
@@ -84,12 +88,14 @@
PIDE/document.scala
PIDE/document_id.scala
PIDE/editor.scala
+ PIDE/line.scala
PIDE/markup.scala
PIDE/markup_tree.scala
PIDE/protocol.scala
PIDE/protocol_message.scala
PIDE/prover.scala
PIDE/query_operation.scala
+ PIDE/rendering.scala
PIDE/resources.scala
PIDE/session.scala
PIDE/text.scala
@@ -113,7 +119,6 @@
System/process_result.scala
System/progress.scala
System/system_channel.scala
- System/utf8.scala
Thy/html.scala
Thy/present.scala
Thy/sessions.scala
@@ -140,19 +145,26 @@
library.scala
term.scala
term_xml.scala
- "../Tools/Graphview/graph_file.scala"
- "../Tools/Graphview/graph_panel.scala"
- "../Tools/Graphview/graphview.scala"
- "../Tools/Graphview/layout.scala"
- "../Tools/Graphview/main_panel.scala"
- "../Tools/Graphview/metrics.scala"
- "../Tools/Graphview/model.scala"
- "../Tools/Graphview/mutator.scala"
- "../Tools/Graphview/mutator_dialog.scala"
- "../Tools/Graphview/mutator_event.scala"
- "../Tools/Graphview/popups.scala"
- "../Tools/Graphview/shapes.scala"
- "../Tools/Graphview/tree_panel.scala"
+ ../Tools/Graphview/graph_file.scala
+ ../Tools/Graphview/graph_panel.scala
+ ../Tools/Graphview/graphview.scala
+ ../Tools/Graphview/layout.scala
+ ../Tools/Graphview/main_panel.scala
+ ../Tools/Graphview/metrics.scala
+ ../Tools/Graphview/model.scala
+ ../Tools/Graphview/mutator.scala
+ ../Tools/Graphview/mutator_dialog.scala
+ ../Tools/Graphview/mutator_event.scala
+ ../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/grammar.scala
+ ../Tools/VSCode/src/protocol.scala
+ ../Tools/VSCode/src/server.scala
+ ../Tools/VSCode/src/vscode_rendering.scala
+ ../Tools/VSCode/src/vscode_resources.scala
)
--- a/src/Pure/conjunction.ML Wed Jan 04 14:26:08 2017 +0100
+++ b/src/Pure/conjunction.ML Wed Jan 04 14:26:19 2017 +0100
@@ -85,13 +85,13 @@
in
val conjunctionD1 =
- Drule.store_standard_thm (Binding.make ("conjunctionD1", @{here})) (conjunctionD #1);
+ Drule.store_standard_thm (Binding.make ("conjunctionD1", \<^here>)) (conjunctionD #1);
val conjunctionD2 =
- Drule.store_standard_thm (Binding.make ("conjunctionD2", @{here})) (conjunctionD #2);
+ Drule.store_standard_thm (Binding.make ("conjunctionD2", \<^here>)) (conjunctionD #2);
val conjunctionI =
- Drule.store_standard_thm (Binding.make ("conjunctionI", @{here}))
+ Drule.store_standard_thm (Binding.make ("conjunctionI", \<^here>))
(Drule.implies_intr_list [A, B]
(Thm.equal_elim
(Thm.symmetric conjunction_def)
--- a/src/Pure/drule.ML Wed Jan 04 14:26:08 2017 +0100
+++ b/src/Pure/drule.ML Wed Jan 04 14:26:19 2017 +0100
@@ -363,13 +363,13 @@
val reflexive_thm =
let val cx = certify (Var(("x",0),TVar(("'a",0),[])))
- in store_standard_thm_open (Binding.make ("reflexive", @{here})) (Thm.reflexive cx) end;
+ in store_standard_thm_open (Binding.make ("reflexive", \<^here>)) (Thm.reflexive cx) end;
val symmetric_thm =
let
val xy = read_prop "x::'a == y::'a";
val thm = Thm.implies_intr xy (Thm.symmetric (Thm.assume xy));
- in store_standard_thm_open (Binding.make ("symmetric", @{here})) thm end;
+ in store_standard_thm_open (Binding.make ("symmetric", \<^here>)) thm end;
val transitive_thm =
let
@@ -378,7 +378,7 @@
val xythm = Thm.assume xy;
val yzthm = Thm.assume yz;
val thm = Thm.implies_intr yz (Thm.transitive xythm yzthm);
- in store_standard_thm_open (Binding.make ("transitive", @{here})) thm end;
+ in store_standard_thm_open (Binding.make ("transitive", \<^here>)) thm end;
fun extensional eq =
let val eq' =
@@ -386,7 +386,7 @@
in Thm.equal_elim (Thm.eta_conversion (Thm.cprop_of eq')) eq' end;
val equals_cong =
- store_standard_thm_open (Binding.make ("equals_cong", @{here}))
+ store_standard_thm_open (Binding.make ("equals_cong", \<^here>))
(Thm.reflexive (read_prop "x::'a == y::'a"));
val imp_cong =
@@ -396,7 +396,7 @@
val AC = read_prop "A ==> C"
val A = read_prop "A"
in
- store_standard_thm_open (Binding.make ("imp_cong", @{here}))
+ store_standard_thm_open (Binding.make ("imp_cong", \<^here>))
(Thm.implies_intr ABC (Thm.equal_intr
(Thm.implies_intr AB (Thm.implies_intr A
(Thm.equal_elim (Thm.implies_elim (Thm.assume ABC) (Thm.assume A))
@@ -413,7 +413,7 @@
val A = read_prop "A"
val B = read_prop "B"
in
- store_standard_thm_open (Binding.make ("swap_prems_eq", @{here}))
+ store_standard_thm_open (Binding.make ("swap_prems_eq", \<^here>))
(Thm.equal_intr
(Thm.implies_intr ABC (Thm.implies_intr B (Thm.implies_intr A
(Thm.implies_elim (Thm.implies_elim (Thm.assume ABC) (Thm.assume A)) (Thm.assume B)))))
@@ -473,12 +473,12 @@
(*The rule V/V, obtains assumption solving for eresolve_tac*)
val asm_rl =
- store_standard_thm_open (Binding.make ("asm_rl", @{here}))
+ store_standard_thm_open (Binding.make ("asm_rl", \<^here>))
(Thm.trivial (read_prop "?psi"));
(*Meta-level cut rule: [| V==>W; V |] ==> W *)
val cut_rl =
- store_standard_thm_open (Binding.make ("cut_rl", @{here}))
+ store_standard_thm_open (Binding.make ("cut_rl", \<^here>))
(Thm.trivial (read_prop "?psi ==> ?theta"));
(*Generalized elim rule for one conclusion; cut_rl with reversed premises:
@@ -488,7 +488,7 @@
val V = read_prop "V";
val VW = read_prop "V ==> W";
in
- store_standard_thm_open (Binding.make ("revcut_rl", @{here}))
+ store_standard_thm_open (Binding.make ("revcut_rl", \<^here>))
(Thm.implies_intr V
(Thm.implies_intr VW (Thm.implies_elim (Thm.assume VW) (Thm.assume V))))
end;
@@ -499,7 +499,7 @@
val V = read_prop "V";
val W = read_prop "W";
val thm = Thm.implies_intr V (Thm.implies_intr W (Thm.assume W));
- in store_standard_thm_open (Binding.make ("thin_rl", @{here})) thm end;
+ in store_standard_thm_open (Binding.make ("thin_rl", \<^here>)) thm end;
(* (!!x. PROP ?V) == PROP ?V Allows removal of redundant parameters*)
val triv_forall_equality =
@@ -508,7 +508,7 @@
val QV = read_prop "!!x::'a. V";
val x = certify (Free ("x", Term.aT []));
in
- store_standard_thm_open (Binding.make ("triv_forall_equality", @{here}))
+ store_standard_thm_open (Binding.make ("triv_forall_equality", \<^here>))
(Thm.equal_intr (Thm.implies_intr QV (Thm.forall_elim x (Thm.assume QV)))
(Thm.implies_intr V (Thm.forall_intr x (Thm.assume V))))
end;
@@ -521,7 +521,7 @@
val AAB = read_prop "Phi ==> Phi ==> Psi";
val A = read_prop "Phi";
in
- store_standard_thm_open (Binding.make ("distinct_prems_rl", @{here}))
+ store_standard_thm_open (Binding.make ("distinct_prems_rl", \<^here>))
(implies_intr_list [AAB, A]
(implies_elim_list (Thm.assume AAB) [Thm.assume A, Thm.assume A]))
end;
@@ -535,7 +535,7 @@
val PQ = read_prop "phi ==> psi";
val QP = read_prop "psi ==> phi";
in
- store_standard_thm_open (Binding.make ("equal_intr_rule", @{here}))
+ store_standard_thm_open (Binding.make ("equal_intr_rule", \<^here>))
(Thm.implies_intr PQ
(Thm.implies_intr QP (Thm.equal_intr (Thm.assume PQ) (Thm.assume QP))))
end;
@@ -546,13 +546,13 @@
val eq = read_prop "phi::prop == psi::prop";
val P = read_prop "phi";
in
- store_standard_thm_open (Binding.make ("equal_elim_rule1", @{here}))
+ store_standard_thm_open (Binding.make ("equal_elim_rule1", \<^here>))
(Thm.equal_elim (Thm.assume eq) (Thm.assume P) |> implies_intr_list [eq, P])
end;
(* PROP ?psi == PROP ?phi ==> PROP ?phi ==> PROP ?psi *)
val equal_elim_rule2 =
- store_standard_thm_open (Binding.make ("equal_elim_rule2", @{here}))
+ store_standard_thm_open (Binding.make ("equal_elim_rule2", \<^here>))
(symmetric_thm RS equal_elim_rule1);
(* PROP ?phi ==> PROP ?phi ==> PROP ?psi ==> PROP ?psi *)
@@ -561,7 +561,7 @@
val P = read_prop "phi";
val Q = read_prop "psi";
val thm = implies_intr_list [P, P, Q] (Thm.assume Q);
- in store_standard_thm_open (Binding.make ("remdups_rl", @{here})) thm end;
+ in store_standard_thm_open (Binding.make ("remdups_rl", \<^here>)) thm end;
@@ -583,15 +583,15 @@
val protect = Thm.apply (certify Logic.protectC);
val protectI =
- store_standard_thm (Binding.concealed (Binding.make ("protectI", @{here})))
+ store_standard_thm (Binding.concealed (Binding.make ("protectI", \<^here>)))
(Thm.equal_elim (Thm.symmetric prop_def) (Thm.assume A));
val protectD =
- store_standard_thm (Binding.concealed (Binding.make ("protectD", @{here})))
+ store_standard_thm (Binding.concealed (Binding.make ("protectD", \<^here>)))
(Thm.equal_elim prop_def (Thm.assume (protect A)));
val protect_cong =
- store_standard_thm_open (Binding.make ("protect_cong", @{here}))
+ store_standard_thm_open (Binding.make ("protect_cong", \<^here>))
(Thm.reflexive (protect A));
fun implies_intr_protected asms th =
@@ -606,7 +606,7 @@
(* term *)
val termI =
- store_standard_thm (Binding.concealed (Binding.make ("termI", @{here})))
+ store_standard_thm (Binding.concealed (Binding.make ("termI", \<^here>)))
(Thm.equal_elim (Thm.symmetric term_def) (Thm.forall_intr A (Thm.trivial A)));
fun mk_term ct =
@@ -637,11 +637,11 @@
| is_sort_constraint _ = false;
val sort_constraintI =
- store_standard_thm (Binding.concealed (Binding.make ("sort_constraintI", @{here})))
+ store_standard_thm (Binding.concealed (Binding.make ("sort_constraintI", \<^here>)))
(Thm.equal_elim (Thm.symmetric sort_constraint_def) (mk_term T));
val sort_constraint_eq =
- store_standard_thm (Binding.concealed (Binding.make ("sort_constraint_eq", @{here})))
+ store_standard_thm (Binding.concealed (Binding.make ("sort_constraint_eq", \<^here>)))
(Thm.equal_intr
(Thm.implies_intr CA (Thm.implies_elim (Thm.assume CA)
(Thm.unvarify_global (Context.the_global_context ()) sort_constraintI)))
@@ -676,7 +676,7 @@
|> Thm.forall_intr cx
|> Thm.implies_intr cphi
|> Thm.implies_intr rhs)
- |> store_standard_thm_open (Binding.make ("norm_hhf_eq", @{here}))
+ |> store_standard_thm_open (Binding.make ("norm_hhf_eq", \<^here>))
end;
val norm_hhf_prop = Logic.dest_equals (Thm.prop_of norm_hhf_eq);
--- a/src/Pure/goal.ML Wed Jan 04 14:26:08 2017 +0100
+++ b/src/Pure/goal.ML Wed Jan 04 14:26:19 2017 +0100
@@ -232,7 +232,7 @@
(Thm.term_of stmt);
in
res
- |> length props > 1 ? Thm.close_derivation
+ |> Thm.close_derivation
|> Conjunction.elim_balanced (length props)
|> map (Assumption.export false ctxt' ctxt)
|> Variable.export ctxt' ctxt
@@ -255,7 +255,7 @@
(* skip proofs *)
-val quick_and_dirty_raw = Config.declare_option ("quick_and_dirty", @{here});
+val quick_and_dirty_raw = Config.declare_option ("quick_and_dirty", \<^here>);
val quick_and_dirty = Config.bool quick_and_dirty_raw;
fun prove_sorry ctxt xs asms prop tac =
--- a/src/Pure/goal_display.ML Wed Jan 04 14:26:08 2017 +0100
+++ b/src/Pure/goal_display.ML Wed Jan 04 14:26:19 2017 +0100
@@ -19,10 +19,10 @@
structure Goal_Display: GOAL_DISPLAY =
struct
-val goals_limit_raw = Config.declare_option ("goals_limit", @{here});
+val goals_limit_raw = Config.declare_option ("goals_limit", \<^here>);
val goals_limit = Config.int goals_limit_raw;
-val show_main_goal_raw = Config.declare_option ("show_main_goal", @{here});
+val show_main_goal_raw = Config.declare_option ("show_main_goal", \<^here>);
val show_main_goal = Config.bool show_main_goal_raw;
--- a/src/Pure/more_thm.ML Wed Jan 04 14:26:08 2017 +0100
+++ b/src/Pure/more_thm.ML Wed Jan 04 14:26:19 2017 +0100
@@ -111,7 +111,7 @@
val untag: string -> attribute
val kind: string -> attribute
val register_proofs: thm list -> theory -> theory
- val join_theory_proofs: theory -> unit
+ val consolidate_theory: theory -> unit
val show_consts_raw: Config.raw
val show_consts: bool Config.T
val show_hyps_raw: Config.raw
@@ -452,8 +452,7 @@
(* close_derivation *)
fun close_derivation thm =
- if Thm.derivation_name thm = "" then Thm.name_derivation "" thm
- else thm;
+ if Thm.derivation_closed thm then thm else Thm.name_derivation "" thm;
(* user renaming of parameters in a subgoal *)
@@ -645,8 +644,8 @@
fun register_proofs more_thms =
Proofs.map (fold (cons o Thm.trim_context) more_thms);
-fun join_theory_proofs thy =
- Thm.join_proofs (map (Thm.transfer thy) (rev (Proofs.get thy)));
+fun consolidate_theory thy =
+ List.app (Thm.consolidate o Thm.transfer thy) (rev (Proofs.get thy));
@@ -654,13 +653,13 @@
(* options *)
-val show_consts_raw = Config.declare_option ("show_consts", @{here});
+val show_consts_raw = Config.declare_option ("show_consts", \<^here>);
val show_consts = Config.bool show_consts_raw;
-val show_hyps_raw = Config.declare ("show_hyps", @{here}) (fn _ => Config.Bool false);
+val show_hyps_raw = Config.declare ("show_hyps", \<^here>) (fn _ => Config.Bool false);
val show_hyps = Config.bool show_hyps_raw;
-val show_tags_raw = Config.declare ("show_tags", @{here}) (fn _ => Config.Bool false);
+val show_tags_raw = Config.declare ("show_tags", \<^here>) (fn _ => Config.Bool false);
val show_tags = Config.bool show_tags_raw;
--- a/src/Pure/pattern.ML Wed Jan 04 14:26:08 2017 +0100
+++ b/src/Pure/pattern.ML Wed Jan 04 14:26:19 2017 +0100
@@ -31,7 +31,7 @@
exception Pattern;
val unify_trace_failure_raw =
- Config.declare ("unify_trace_failure", @{here}) (fn _ => Config.Bool false);
+ Config.declare ("unify_trace_failure", \<^here>) (fn _ => Config.Bool false);
val unify_trace_failure = Config.bool unify_trace_failure_raw;
fun string_of_term ctxt env binders t =
--- a/src/Pure/primitive_defs.ML Wed Jan 04 14:26:08 2017 +0100
+++ b/src/Pure/primitive_defs.ML Wed Jan 04 14:26:19 2017 +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 [];
--- a/src/Pure/proofterm.ML Wed Jan 04 14:26:08 2017 +0100
+++ b/src/Pure/proofterm.ML Wed Jan 04 14:26:19 2017 +0100
@@ -10,6 +10,7 @@
sig
val proofs: int Unsynchronized.ref
+ type thm_node
datatype proof =
MinProof
| PBound of int
@@ -25,7 +26,7 @@
| PThm of serial * ((string * term * typ list option) * proof_body future)
and proof_body = PBody of
{oracles: (string * term) Ord_List.T,
- thms: (serial * (string * term * proof_body future)) Ord_List.T,
+ thms: (serial * thm_node) Ord_List.T,
proof: proof}
val %> : proof * term -> proof
@@ -35,13 +36,17 @@
sig
include BASIC_PROOFTERM
+ type pthm = serial * thm_node
type oracle = string * term
- type pthm = serial * (string * term * proof_body future)
val proof_of: proof_body -> proof
+ val thm_node_name: thm_node -> string
+ val thm_node_prop: thm_node -> term
+ val thm_node_body: thm_node -> proof_body future
val join_proof: proof_body future -> proof
val fold_proof_atoms: bool -> (proof -> 'a -> 'a) -> proof list -> 'a -> 'a
- val fold_body_thms: (string * term * proof_body -> 'a -> 'a) -> proof_body list -> 'a -> 'a
- val join_bodies: proof_body list -> unit
+ val fold_body_thms: ({serial: serial, name: string, prop: term, body: proof_body} -> 'a -> 'a) ->
+ proof_body list -> 'a -> 'a
+ val consolidate: proof_body -> unit
val peek_status: proof_body list -> {failed: bool, oracle: bool, unfinished: bool}
val oracle_ord: oracle * oracle -> order
@@ -60,6 +65,8 @@
(** primitive operations **)
val proofs_enabled: unit -> bool
+ val atomic_proof: proof -> bool
+ val compact_proof: proof -> bool
val proof_combt: proof * term list -> proof
val proof_combt': proof * term option list -> proof
val proof_combP: proof * proof list -> proof
@@ -173,16 +180,35 @@
| PThm of serial * ((string * term * typ list option) * proof_body future)
and proof_body = PBody of
{oracles: (string * term) Ord_List.T,
- thms: (serial * (string * term * proof_body future)) Ord_List.T,
- proof: proof};
+ thms: (serial * thm_node) Ord_List.T,
+ proof: proof}
+and thm_node =
+ Thm_Node of {name: string, prop: term, body: proof_body future, consolidate: unit lazy};
type oracle = string * term;
-type pthm = serial * (string * term * proof_body future);
+type pthm = serial * thm_node;
fun proof_of (PBody {proof, ...}) = proof;
val join_proof = Future.join #> proof_of;
-fun join_thms (thms: pthm list) = ignore (Future.joins (map (#3 o #2) thms));
+fun rep_thm_node (Thm_Node args) = args;
+val thm_node_name = #name o rep_thm_node;
+val thm_node_prop = #prop o rep_thm_node;
+val thm_node_body = #body o rep_thm_node;
+val thm_node_consolidate = #consolidate o rep_thm_node;
+
+fun join_thms (thms: pthm list) =
+ Future.joins (map (thm_node_body o #2) thms);
+
+fun consolidate (PBody {thms, ...}) =
+ List.app (Lazy.force o thm_node_consolidate o #2) thms;
+
+fun make_thm_node name prop body =
+ Thm_Node {name = name, prop = prop, body = body,
+ consolidate =
+ Lazy.lazy (fn () =>
+ let val PBody {thms, ...} = Future.join body
+ in List.app consolidate (join_thms thms) end)};
(***** proof atoms *****)
@@ -205,27 +231,27 @@
fun fold_body_thms f =
let
fun app (PBody {thms, ...}) =
- tap join_thms thms |> fold (fn (i, (name, prop, body)) => fn (x, seen) =>
+ tap join_thms thms |> fold (fn (i, thm_node) => fn (x, seen) =>
if Inttab.defined seen i then (x, seen)
else
let
- val body' = Future.join body;
- val (x', seen') = app body' (x, Inttab.update (i, ()) seen);
- in (f (name, prop, body') x', seen') end);
+ val name = thm_node_name thm_node;
+ val prop = thm_node_prop thm_node;
+ val body = Future.join (thm_node_body thm_node);
+ val (x', seen') = app body (x, Inttab.update (i, ()) seen);
+ in (f {serial = i, name = name, prop = prop, body = body} x', seen') end);
in fn bodies => fn x => #1 (fold app bodies (x, Inttab.empty)) end;
-fun join_bodies bodies = fold_body_thms (fn _ => fn () => ()) bodies ();
-
fun peek_status bodies =
let
fun status (PBody {oracles, thms, ...}) x =
let
val ((oracle, unfinished, failed), seen) =
- (thms, x) |-> fold (fn (i, (_, _, body)) => fn (st, seen) =>
+ (thms, x) |-> fold (fn (i, thm_node) => fn (st, seen) =>
if Inttab.defined seen i then (st, seen)
else
let val seen' = Inttab.update (i, ()) seen in
- (case Future.peek body of
+ (case Future.peek (thm_node_body thm_node) of
SOME (Exn.Res body') => status body' (st, seen')
| SOME (Exn.Exn _) =>
let val (oracle, unfinished, _) = st
@@ -243,7 +269,7 @@
(* proof body *)
val oracle_ord = prod_ord fast_string_ord Term_Ord.fast_term_ord;
-fun thm_ord ((i, _): pthm, (j, _)) = int_ord (j, i);
+fun thm_ord ((i, _): pthm, (j, _): pthm) = int_ord (j, i);
val unions_oracles = Ord_List.unions oracle_ord;
val unions_thms = Ord_List.unions thm_ord;
@@ -251,12 +277,12 @@
val all_oracles_of =
let
fun collect (PBody {oracles, thms, ...}) =
- tap join_thms thms |> fold (fn (i, (_, _, body)) => fn (x, seen) =>
+ tap join_thms thms |> fold (fn (i, thm_node) => fn (x, seen) =>
if Inttab.defined seen i then (x, seen)
else
let
- val body' = Future.join body;
- val (x', seen') = collect body' (x, Inttab.update (i, ()) seen);
+ val body = Future.join (thm_node_body thm_node);
+ val (x', seen') = collect body (x, Inttab.update (i, ()) seen);
in (if null oracles then x' else oracles :: x', seen') end);
in fn body => unions_oracles (#1 (collect body ([], Inttab.empty))) end;
@@ -264,7 +290,7 @@
let
val (oracles, thms) = fold_proof_atoms false
(fn Oracle (s, prop, _) => apfst (cons (s, prop))
- | PThm (i, ((name, prop, _), body)) => apsnd (cons (i, (name, prop, body)))
+ | PThm (i, ((name, prop, _), body)) => apsnd (cons (i, make_thm_node name prop body))
| _ => I) [prf] ([], []);
in
PBody
@@ -308,8 +334,9 @@
([int_atom a, b], triple term (option (list typ)) proof_body (c, d, Future.join body))]
and proof_body (PBody {oracles, thms, proof = prf}) =
triple (list (pair string term)) (list pthm) proof (oracles, thms, prf)
-and pthm (a, (b, c, body)) =
- pair int (triple string term proof_body) (a, (b, c, Future.join body));
+and pthm (a, thm_node) =
+ pair int (triple string term proof_body)
+ (a, (thm_node_name thm_node, thm_node_prop thm_node, Future.join (thm_node_body thm_node)));
in
@@ -345,7 +372,7 @@
in PBody {oracles = a, thms = b, proof = c} end
and pthm x =
let val (a, (b, c, d)) = pair int (triple string term proof_body) x
- in (a, (b, c, Future.value d)) end;
+ in (a, make_thm_node b c (Future.value d)) end;
in
@@ -357,6 +384,19 @@
(***** proof objects with different levels of detail *****)
+fun atomic_proof prf =
+ (case prf of
+ Abst _ => false
+ | AbsP _ => false
+ | op % _ => false
+ | op %% _ => false
+ | Promise _ => false
+ | _ => true);
+
+fun compact_proof (prf % _) = compact_proof prf
+ | compact_proof (prf1 %% prf2) = atomic_proof prf2 andalso compact_proof prf1
+ | compact_proof prf = atomic_proof prf;
+
fun (prf %> t) = prf % SOME t;
val proof_combt = Library.foldl (op %>);
@@ -533,11 +573,13 @@
prf_add_loose_bnos plev tlev prf2
(prf_add_loose_bnos plev tlev prf1 p)
| prf_add_loose_bnos plev tlev (prf % opt) (is, js) =
- prf_add_loose_bnos plev tlev prf (case opt of
+ prf_add_loose_bnos plev tlev prf
+ (case opt of
NONE => (is, insert (op =) ~1 js)
| SOME t => (is, add_loose_bnos (t, tlev, js)))
| prf_add_loose_bnos plev tlev (AbsP (_, opt, prf)) (is, js) =
- prf_add_loose_bnos (plev+1) tlev prf (case opt of
+ prf_add_loose_bnos (plev+1) tlev prf
+ (case opt of
NONE => (is, insert (op =) ~1 js)
| SOME t => (is, add_loose_bnos (t, tlev, js)))
| prf_add_loose_bnos plev tlev (Abst (_, _, prf)) p =
@@ -635,7 +677,7 @@
handle Same.SAME => prf % Same.map_option (subst' lev) t)
| subst _ _ = raise Same.SAME
and substh lev prf = (subst lev prf handle Same.SAME => prf);
- in case args of [] => prf | _ => substh 0 prf end;
+ in (case args of [] => prf | _ => substh 0 prf) end;
fun prf_subst_pbounds args prf =
let
@@ -651,7 +693,7 @@
| subst (prf % t) Plev tlev = subst prf Plev tlev % t
| subst _ _ _ = raise Same.SAME
and substh prf Plev tlev = (subst prf Plev tlev handle Same.SAME => prf)
- in case args of [] => prf | _ => substh prf 0 0 end;
+ in (case args of [] => prf | _ => substh prf 0 0) end;
(**** Freezing and thawing of variables in proof terms ****)
@@ -992,7 +1034,7 @@
| Var _ => SOME (remove_types t)
| _ => NONE) %
(case head_of g of
- Abs _ => SOME (remove_types u)
+ Abs _ => SOME (remove_types u)
| Var _ => SOME (remove_types u)
| _ => NONE) %% prf1 %% prf2
| _ => prf % NONE % NONE %% prf1 %% prf2)
@@ -1105,7 +1147,8 @@
add_npvars q p Ts (vs, if p andalso q then betapply (t, Var (("",0), T)) else t)
| add_npvars q p Ts (vs, Abs (_, T, t)) = add_npvars q p (T::Ts) (vs, t)
| add_npvars _ _ Ts (vs, t) = add_npvars' Ts (vs, t)
-and add_npvars' Ts (vs, t) = (case strip_comb t of
+and add_npvars' Ts (vs, t) =
+ (case strip_comb t of
(Var (ixn, _), ts) => if test_args [] ts then vs
else Library.foldl (add_npvars' Ts)
(AList.update (op =) (ixn,
@@ -1115,19 +1158,20 @@
fun prop_vars (Const ("Pure.imp", _) $ P $ Q) = union (op =) (prop_vars P) (prop_vars Q)
| prop_vars (Const ("Pure.all", _) $ Abs (_, _, t)) = prop_vars t
- | prop_vars t = (case strip_comb t of
- (Var (ixn, _), _) => [ixn] | _ => []);
+ | prop_vars t = (case strip_comb t of (Var (ixn, _), _) => [ixn] | _ => []);
fun is_proj t =
let
- fun is_p i t = (case strip_comb t of
+ fun is_p i t =
+ (case strip_comb t of
(Bound _, []) => false
| (Bound j, ts) => j >= i orelse exists (is_p i) ts
| (Abs (_, _, u), _) => is_p (i+1) u
| (_, ts) => exists (is_p i) ts)
- in (case strip_abs_body t of
- Bound _ => true
- | t' => is_p 0 t')
+ in
+ (case strip_abs_body t of
+ Bound _ => true
+ | t' => is_p 0 t')
end;
fun needed_vars prop =
@@ -1196,7 +1240,8 @@
val (ts', ts'') = chop (length vs) ts;
val insts = take (length ts') (map (fst o dest_Var) vs) ~~ ts';
val nvs = Library.foldl (fn (ixns', (ixn, ixns)) =>
- insert (op =) ixn (case AList.lookup (op =) insts ixn of
+ insert (op =) ixn
+ (case AList.lookup (op =) insts ixn of
SOME (SOME t) => if is_proj t then union (op =) ixns ixns' else ixns'
| _ => union (op =) ixns ixns'))
(needed prop ts'' prfs, add_npvars false true [] ([], prop));
@@ -1250,12 +1295,12 @@
fun mtch Ts i j (pinst, tinst) (Hyp (Var (ixn, _)), prf) =
if i = 0 andalso j = 0 then ((ixn, prf) :: pinst, tinst)
- else (case apfst (flt i) (apsnd (flt j)
- (prf_add_loose_bnos 0 0 prf ([], []))) of
+ else
+ (case apfst (flt i) (apsnd (flt j) (prf_add_loose_bnos 0 0 prf ([], []))) of
([], []) => ((ixn, incr_pboundvars (~i) (~j) prf) :: pinst, tinst)
- | ([], _) => if j = 0 then
- ((ixn, incr_pboundvars (~i) (~j) prf) :: pinst, tinst)
- else raise PMatch
+ | ([], _) =>
+ if j = 0 then ((ixn, incr_pboundvars (~i) (~j) prf) :: pinst, tinst)
+ else raise PMatch
| _ => raise PMatch)
| mtch Ts i j inst (prf1 % opt1, prf2 % opt2) =
optmatch (matcht Ts j) (mtch Ts i j inst (prf1, prf2)) (opt1, opt2)
@@ -1389,45 +1434,57 @@
| rew0 Ts hs prf = rew Ts hs prf;
fun rew1 _ _ (Hyp (Var _)) _ = NONE
- | rew1 Ts hs skel prf = (case rew2 Ts hs skel prf of
- SOME prf1 => (case rew0 Ts hs prf1 of
- SOME (prf2, skel') => SOME (the_default prf2 (rew1 Ts hs skel' prf2))
- | NONE => SOME prf1)
- | NONE => (case rew0 Ts hs prf of
- SOME (prf1, skel') => SOME (the_default prf1 (rew1 Ts hs skel' prf1))
- | NONE => NONE))
+ | rew1 Ts hs skel prf =
+ (case rew2 Ts hs skel prf of
+ SOME prf1 =>
+ (case rew0 Ts hs prf1 of
+ SOME (prf2, skel') => SOME (the_default prf2 (rew1 Ts hs skel' prf2))
+ | NONE => SOME prf1)
+ | NONE =>
+ (case rew0 Ts hs prf of
+ SOME (prf1, skel') => SOME (the_default prf1 (rew1 Ts hs skel' prf1))
+ | NONE => NONE))
- and rew2 Ts hs skel (prf % SOME t) = (case prf of
+ and rew2 Ts hs skel (prf % SOME t) =
+ (case prf of
Abst (_, _, body) =>
let val prf' = prf_subst_bounds [t] body
in SOME (the_default prf' (rew2 Ts hs no_skel prf')) end
- | _ => (case rew1 Ts hs (case skel of skel' % _ => skel' | _ => no_skel) prf of
- SOME prf' => SOME (prf' % SOME t)
- | NONE => NONE))
+ | _ =>
+ (case rew1 Ts hs (case skel of skel' % _ => skel' | _ => no_skel) prf of
+ SOME prf' => SOME (prf' % SOME t)
+ | NONE => NONE))
| rew2 Ts hs skel (prf % NONE) = Option.map (fn prf' => prf' % NONE)
(rew1 Ts hs (case skel of skel' % _ => skel' | _ => no_skel) prf)
- | rew2 Ts hs skel (prf1 %% prf2) = (case prf1 of
+ | rew2 Ts hs skel (prf1 %% prf2) =
+ (case prf1 of
AbsP (_, _, body) =>
let val prf' = prf_subst_pbounds [prf2] body
in SOME (the_default prf' (rew2 Ts hs no_skel prf')) end
| _ =>
- let val (skel1, skel2) = (case skel of
- skel1 %% skel2 => (skel1, skel2)
- | _ => (no_skel, no_skel))
- in case rew1 Ts hs skel1 prf1 of
- SOME prf1' => (case rew1 Ts hs skel2 prf2 of
+ let
+ val (skel1, skel2) =
+ (case skel of
+ skel1 %% skel2 => (skel1, skel2)
+ | _ => (no_skel, no_skel))
+ in
+ (case rew1 Ts hs skel1 prf1 of
+ SOME prf1' =>
+ (case rew1 Ts hs skel2 prf2 of
SOME prf2' => SOME (prf1' %% prf2')
| NONE => SOME (prf1' %% prf2))
- | NONE => (case rew1 Ts hs skel2 prf2 of
+ | NONE =>
+ (case rew1 Ts hs skel2 prf2 of
SOME prf2' => SOME (prf1 %% prf2')
- | NONE => NONE)
+ | NONE => NONE))
end)
- | rew2 Ts hs skel (Abst (s, T, prf)) = (case rew1 (the_default dummyT T :: Ts) hs
+ | rew2 Ts hs skel (Abst (s, T, prf)) =
+ (case rew1 (the_default dummyT T :: Ts) hs
(case skel of Abst (_, _, skel') => skel' | _ => no_skel) prf of
SOME prf' => SOME (Abst (s, T, prf'))
| NONE => NONE)
- | rew2 Ts hs skel (AbsP (s, t, prf)) = (case rew1 Ts (t :: hs)
- (case skel of AbsP (_, _, skel') => skel' | _ => no_skel) prf of
+ | rew2 Ts hs skel (AbsP (s, t, prf)) =
+ (case rew1 Ts (t :: hs) (case skel of AbsP (_, _, skel') => skel' | _ => no_skel) prf of
SOME prf' => SOME (AbsP (s, t, prf'))
| NONE => NONE)
| rew2 _ _ _ _ = NONE;
@@ -1476,6 +1533,8 @@
fun fulfill_norm_proof thy ps body0 =
let
+ val _ = List.app (consolidate o #2) ps;
+ val _ = consolidate body0;
val PBody {oracles = oracles0, thms = thms0, proof = proof0} = body0;
val oracles =
unions_oracles
@@ -1573,15 +1632,14 @@
else new_prf ()
| _ => new_prf ());
val head = PThm (i, ((name, prop1, NONE), body'));
- in ((i, (name, prop1, body')), head, args, argsP, args1) end;
+ in ((i, make_thm_node name prop1 body'), head, args, argsP, args1) end;
fun thm_proof thy name shyps hyps concl promises body =
let val (pthm, head, args, argsP, _) = prepare_thm_proof thy name shyps hyps concl promises body
in (pthm, proof_combP (proof_combt' (head, args), argsP)) end;
fun unconstrain_thm_proof thy shyps concl promises body =
- let
- val (pthm, head, _, _, args) = prepare_thm_proof thy "" shyps [] concl promises body
+ let val (pthm, head, _, _, args) = prepare_thm_proof thy "" shyps [] concl promises body
in (pthm, proof_combt' (head, args)) end;
--- a/src/Pure/pure_syn.ML Wed Jan 04 14:26:08 2017 +0100
+++ b/src/Pure/pure_syn.ML Wed Jan 04 14:26:19 2017 +0100
@@ -16,49 +16,49 @@
val semi = Scan.option (Parse.$$$ ";");
val _ =
- Outer_Syntax.command ("chapter", @{here}) "chapter heading"
+ Outer_Syntax.command ("chapter", \<^here>) "chapter heading"
(Parse.opt_target -- Parse.document_source --| semi
>> Thy_Output.document_command {markdown = false});
val _ =
- Outer_Syntax.command ("section", @{here}) "section heading"
+ Outer_Syntax.command ("section", \<^here>) "section heading"
(Parse.opt_target -- Parse.document_source --| semi
>> Thy_Output.document_command {markdown = false});
val _ =
- Outer_Syntax.command ("subsection", @{here}) "subsection heading"
+ Outer_Syntax.command ("subsection", \<^here>) "subsection heading"
(Parse.opt_target -- Parse.document_source --| semi
>> Thy_Output.document_command {markdown = false});
val _ =
- Outer_Syntax.command ("subsubsection", @{here}) "subsubsection heading"
+ Outer_Syntax.command ("subsubsection", \<^here>) "subsubsection heading"
(Parse.opt_target -- Parse.document_source --| semi
>> Thy_Output.document_command {markdown = false});
val _ =
- Outer_Syntax.command ("paragraph", @{here}) "paragraph heading"
+ Outer_Syntax.command ("paragraph", \<^here>) "paragraph heading"
(Parse.opt_target -- Parse.document_source --| semi
>> Thy_Output.document_command {markdown = false});
val _ =
- Outer_Syntax.command ("subparagraph", @{here}) "subparagraph heading"
+ Outer_Syntax.command ("subparagraph", \<^here>) "subparagraph heading"
(Parse.opt_target -- Parse.document_source --| semi
>> Thy_Output.document_command {markdown = false});
val _ =
- Outer_Syntax.command ("text", @{here}) "formal comment (primary style)"
+ Outer_Syntax.command ("text", \<^here>) "formal comment (primary style)"
(Parse.opt_target -- Parse.document_source >> Thy_Output.document_command {markdown = true});
val _ =
- Outer_Syntax.command ("txt", @{here}) "formal comment (secondary style)"
+ Outer_Syntax.command ("txt", \<^here>) "formal comment (secondary style)"
(Parse.opt_target -- Parse.document_source >> Thy_Output.document_command {markdown = true});
val _ =
- Outer_Syntax.command ("text_raw", @{here}) "LaTeX text (without surrounding environment)"
+ Outer_Syntax.command ("text_raw", \<^here>) "LaTeX text (without surrounding environment)"
(Parse.opt_target -- Parse.document_source >> Thy_Output.document_command {markdown = true});
val _ =
- Outer_Syntax.command ("theory", @{here}) "begin theory"
+ Outer_Syntax.command ("theory", \<^here>) "begin theory"
(Thy_Header.args >>
(fn _ => Toplevel.init_theory (fn () => error "Missing theory initialization")));
--- a/src/Pure/pure_thy.ML Wed Jan 04 14:26:08 2017 +0100
+++ b/src/Pure/pure_thy.ML Wed Jan 04 14:26:19 2017 +0100
@@ -67,16 +67,16 @@
(Sign.map_naming (Name_Space.set_theory_name Context.PureN) #>
Old_Appl_Syntax.put false #>
Sign.add_types_global
- [(Binding.make ("fun", @{here}), 2, NoSyn),
- (Binding.make ("prop", @{here}), 0, NoSyn),
- (Binding.make ("itself", @{here}), 1, NoSyn),
- (Binding.make ("dummy", @{here}), 0, NoSyn)]
+ [(Binding.make ("fun", \<^here>), 2, NoSyn),
+ (Binding.make ("prop", \<^here>), 0, NoSyn),
+ (Binding.make ("itself", \<^here>), 1, NoSyn),
+ (Binding.make ("dummy", \<^here>), 0, NoSyn)]
#> Theory.add_deps_global "fun" ((Defs.Type, "fun"), [typ "'a", typ "'b"]) []
#> Theory.add_deps_global "prop" ((Defs.Type, "prop"), []) []
#> Theory.add_deps_global "itself" ((Defs.Type, "itself"), [typ "'a"]) []
#> Theory.add_deps_global "dummy" ((Defs.Type, "dummy"), []) []
#> Sign.add_nonterminals_global
- (map (fn name => Binding.make (name, @{here}))
+ (map (fn name => Binding.make (name, \<^here>))
(Lexicon.terminals @ ["logic", "type", "types", "sort", "classes",
"args", "cargs", "pttrn", "pttrns", "idt", "idts", "aprop", "asms",
"any", "prop'", "num_const", "float_const", "num_position",
@@ -194,12 +194,12 @@
#> Sign.add_syntax ("", false)
[(const "Pure.prop", typ "prop => prop", mixfix ("_", [0], 0))]
#> Sign.add_consts
- [(qualify (Binding.make ("eq", @{here})), typ "'a => 'a => prop", infix_ ("\<equiv>", 2)),
- (qualify (Binding.make ("imp", @{here})), typ "prop => prop => prop", infixr_ ("\<Longrightarrow>", 1)),
- (qualify (Binding.make ("all", @{here})), typ "('a => prop) => prop", binder ("\<And>", 0, 0)),
- (qualify (Binding.make ("prop", @{here})), typ "prop => prop", NoSyn),
- (qualify (Binding.make ("type", @{here})), typ "'a itself", NoSyn),
- (qualify (Binding.make ("dummy_pattern", @{here})), typ "'a", Mixfix.mixfix "'_")]
+ [(qualify (Binding.make ("eq", \<^here>)), typ "'a => 'a => prop", infix_ ("\<equiv>", 2)),
+ (qualify (Binding.make ("imp", \<^here>)), typ "prop => prop => prop", infixr_ ("\<Longrightarrow>", 1)),
+ (qualify (Binding.make ("all", \<^here>)), typ "('a => prop) => prop", binder ("\<And>", 0, 0)),
+ (qualify (Binding.make ("prop", \<^here>)), typ "prop => prop", NoSyn),
+ (qualify (Binding.make ("type", \<^here>)), typ "'a itself", NoSyn),
+ (qualify (Binding.make ("dummy_pattern", \<^here>)), typ "'a", Mixfix.mixfix "'_")]
#> Theory.add_deps_global "Pure.eq" ((Defs.Const, "Pure.eq"), [typ "'a"]) []
#> Theory.add_deps_global "Pure.imp" ((Defs.Const, "Pure.imp"), []) []
#> Theory.add_deps_global "Pure.all" ((Defs.Const, "Pure.all"), [typ "'a"]) []
@@ -209,21 +209,21 @@
#> Sign.parse_translation Syntax_Trans.pure_parse_translation
#> Sign.print_ast_translation Syntax_Trans.pure_print_ast_translation
#> Sign.add_consts
- [(qualify (Binding.make ("term", @{here})), typ "'a => prop", NoSyn),
- (qualify (Binding.make ("sort_constraint", @{here})), typ "'a itself => prop", NoSyn),
- (qualify (Binding.make ("conjunction", @{here})), typ "prop => prop => prop", NoSyn)]
+ [(qualify (Binding.make ("term", \<^here>)), typ "'a => prop", NoSyn),
+ (qualify (Binding.make ("sort_constraint", \<^here>)), typ "'a itself => prop", NoSyn),
+ (qualify (Binding.make ("conjunction", \<^here>)), typ "prop => prop => prop", NoSyn)]
#> Sign.local_path
#> (Global_Theory.add_defs false o map Thm.no_attributes)
- [(Binding.make ("prop_def", @{here}),
+ [(Binding.make ("prop_def", \<^here>),
prop "(CONST Pure.prop :: prop => prop) (A::prop) == A::prop"),
- (Binding.make ("term_def", @{here}),
+ (Binding.make ("term_def", \<^here>),
prop "(CONST Pure.term :: 'a => prop) (x::'a) == (!!A::prop. A ==> A)"),
- (Binding.make ("sort_constraint_def", @{here}),
+ (Binding.make ("sort_constraint_def", \<^here>),
prop "(CONST Pure.sort_constraint :: 'a itself => prop) (CONST Pure.type :: 'a itself) ==\
\ (CONST Pure.term :: 'a itself => prop) (CONST Pure.type :: 'a itself)"),
- (Binding.make ("conjunction_def", @{here}),
+ (Binding.make ("conjunction_def", \<^here>),
prop "(A &&& B) == (!!C::prop. (A ==> B ==> C) ==> C)")] #> snd
#> fold (fn (a, prop) =>
- snd o Thm.add_axiom_global (Binding.make (a, @{here}), prop)) Proofterm.equality_axms);
+ snd o Thm.add_axiom_global (Binding.make (a, \<^here>), prop)) Proofterm.equality_axms);
end;
--- a/src/Pure/raw_simplifier.ML Wed Jan 04 14:26:08 2017 +0100
+++ b/src/Pure/raw_simplifier.ML Wed Jan 04 14:26:19 2017 +0100
@@ -389,11 +389,11 @@
(* simp depth *)
-val simp_depth_limit_raw = Config.declare ("simp_depth_limit", @{here}) (K (Config.Int 100));
+val simp_depth_limit_raw = Config.declare ("simp_depth_limit", \<^here>) (K (Config.Int 100));
val simp_depth_limit = Config.int simp_depth_limit_raw;
val simp_trace_depth_limit_raw =
- Config.declare ("simp_trace_depth_limit", @{here}) (fn _ => Config.Int 1);
+ Config.declare ("simp_trace_depth_limit", \<^here>) (fn _ => Config.Int 1);
val simp_trace_depth_limit = Config.int simp_trace_depth_limit_raw;
fun inc_simp_depth ctxt =
@@ -412,10 +412,10 @@
exception SIMPLIFIER of string * thm list;
-val simp_debug_raw = Config.declare ("simp_debug", @{here}) (K (Config.Bool false));
+val simp_debug_raw = Config.declare ("simp_debug", \<^here>) (K (Config.Bool false));
val simp_debug = Config.bool simp_debug_raw;
-val simp_trace_raw = Config.declare ("simp_trace", @{here}) (fn _ => Config.Bool false);
+val simp_trace_raw = Config.declare ("simp_trace", \<^here>) (fn _ => Config.Bool false);
val simp_trace = Config.bool simp_trace_raw;
fun cond_warning ctxt msg =
--- a/src/Pure/simplifier.ML Wed Jan 04 14:26:08 2017 +0100
+++ b/src/Pure/simplifier.ML Wed Jan 04 14:26:19 2017 +0100
@@ -332,23 +332,23 @@
(** method syntax **)
val cong_modifiers =
- [Args.$$$ congN -- Args.colon >> K (Method.modifier cong_add @{here}),
- Args.$$$ congN -- Args.add -- Args.colon >> K (Method.modifier cong_add @{here}),
- Args.$$$ congN -- Args.del -- Args.colon >> K (Method.modifier cong_del @{here})];
+ [Args.$$$ congN -- Args.colon >> K (Method.modifier cong_add \<^here>),
+ Args.$$$ congN -- Args.add -- Args.colon >> K (Method.modifier cong_add \<^here>),
+ Args.$$$ congN -- Args.del -- Args.colon >> K (Method.modifier cong_del \<^here>)];
val simp_modifiers =
- [Args.$$$ simpN -- Args.colon >> K (Method.modifier simp_add @{here}),
- Args.$$$ simpN -- Args.add -- Args.colon >> K (Method.modifier simp_add @{here}),
- Args.$$$ simpN -- Args.del -- Args.colon >> K (Method.modifier simp_del @{here}),
+ [Args.$$$ simpN -- Args.colon >> K (Method.modifier simp_add \<^here>),
+ Args.$$$ simpN -- Args.add -- Args.colon >> K (Method.modifier simp_add \<^here>),
+ Args.$$$ simpN -- Args.del -- Args.colon >> K (Method.modifier simp_del \<^here>),
Args.$$$ simpN -- Args.$$$ onlyN -- Args.colon >>
- K {init = Raw_Simplifier.clear_simpset, attribute = simp_add, pos = @{here}}]
+ K {init = Raw_Simplifier.clear_simpset, attribute = simp_add, pos = \<^here>}]
@ cong_modifiers;
val simp_modifiers' =
- [Args.add -- Args.colon >> K (Method.modifier simp_add @{here}),
- Args.del -- Args.colon >> K (Method.modifier simp_del @{here}),
+ [Args.add -- Args.colon >> K (Method.modifier simp_add \<^here>),
+ Args.del -- Args.colon >> K (Method.modifier simp_del \<^here>),
Args.$$$ onlyN -- Args.colon >>
- K {init = Raw_Simplifier.clear_simpset, attribute = simp_add, pos = @{here}}]
+ K {init = Raw_Simplifier.clear_simpset, attribute = simp_add, pos = \<^here>}]
@ cong_modifiers;
val simp_options =
--- a/src/Pure/skip_proof.ML Wed Jan 04 14:26:08 2017 +0100
+++ b/src/Pure/skip_proof.ML Wed Jan 04 14:26:19 2017 +0100
@@ -19,7 +19,7 @@
fun report ctxt =
if Context_Position.is_visible ctxt then
- Output.report [Markup.markup Markup.bad "Skipped proof"]
+ Output.report [Markup.markup (Markup.bad ()) "Skipped proof"]
else ();
@@ -27,7 +27,7 @@
val (_, make_thm_cterm) =
Context.>>>
- (Context.map_theory_result (Thm.add_oracle (Binding.make ("skip_proof", @{here}), I)));
+ (Context.map_theory_result (Thm.add_oracle (Binding.make ("skip_proof", \<^here>), I)));
fun make_thm thy prop = make_thm_cterm (Thm.global_cterm_of thy prop);
--- a/src/Pure/thm.ML Wed Jan 04 14:26:08 2017 +0100
+++ b/src/Pure/thm.ML Wed Jan 04 14:26:19 2017 +0100
@@ -86,9 +86,10 @@
val proof_bodies_of: thm list -> proof_body list
val proof_body_of: thm -> proof_body
val proof_of: thm -> proof
- val join_proofs: thm list -> unit
+ val consolidate: thm -> unit
val peek_status: thm -> {oracle: bool, unfinished: bool, failed: bool}
val future: thm future -> cterm -> thm
+ val derivation_closed: thm -> bool
val derivation_name: thm -> string
val name_derivation: string -> thm -> thm
val axiom: theory -> string -> thm
@@ -584,21 +585,14 @@
and join_promises_of thms = join_promises (Ord_List.make promise_ord (maps raw_promises_of thms));
fun fulfill_body (th as Thm (Deriv {promises, body}, _)) =
- Proofterm.fulfill_norm_proof (theory_of_thm th) (fulfill_promises promises) body
-and fulfill_promises promises =
- map fst promises ~~ map fulfill_body (Future.joins (map snd promises));
+ let val fulfilled_promises = map #1 promises ~~ map fulfill_body (Future.joins (map #2 promises))
+ in Proofterm.fulfill_norm_proof (theory_of_thm th) fulfilled_promises body end;
-fun proof_bodies_of thms =
- let
- val _ = join_promises_of thms;
- val bodies = map fulfill_body thms;
- val _ = Proofterm.join_bodies bodies;
- in bodies end;
-
+fun proof_bodies_of thms = (join_promises_of thms; map fulfill_body thms);
val proof_body_of = singleton proof_bodies_of;
val proof_of = Proofterm.proof_of o proof_body_of;
-val join_proofs = ignore o proof_bodies_of;
+val consolidate = ignore o proof_bodies_of o single;
(* derivation status *)
@@ -655,6 +649,10 @@
(* closed derivations with official name *)
(*non-deterministic, depends on unknown promises*)
+fun derivation_closed (Thm (Deriv {body, ...}, _)) =
+ Proofterm.compact_proof (Proofterm.proof_of body);
+
+(*non-deterministic, depends on unknown promises*)
fun derivation_name (Thm (Deriv {body, ...}, {shyps, hyps, prop, ...})) =
Proofterm.get_name shyps hyps prop (Proofterm.proof_of body);
@@ -1305,9 +1303,9 @@
val _ = null tfrees orelse err ("illegal free type variables " ^ commas_quote tfrees);
val ps = map (apsnd (Future.map fulfill_body)) promises;
- val (pthm as (_, (_, prop', _)), proof) =
- Proofterm.unconstrain_thm_proof thy shyps prop ps body;
+ val (pthm, proof) = Proofterm.unconstrain_thm_proof thy shyps prop ps body;
val der' = make_deriv [] [] [pthm] proof;
+ val prop' = Proofterm.thm_node_prop (#2 pthm);
in
Thm (der',
{cert = cert,
--- a/src/Pure/type_infer.ML Wed Jan 04 14:26:08 2017 +0100
+++ b/src/Pure/type_infer.ML Wed Jan 04 14:26:19 2017 +0100
@@ -101,7 +101,7 @@
(* fixate -- introduce fresh type variables *)
val object_logic =
- Config.bool (Config.declare ("Type_Infer.object_logic", @{here}) (K (Config.Bool true)));
+ Config.bool (Config.declare ("Type_Infer.object_logic", \<^here>) (K (Config.Bool true)));
fun fixate ctxt pattern ts =
let
--- a/src/Pure/type_infer_context.ML Wed Jan 04 14:26:08 2017 +0100
+++ b/src/Pure/type_infer_context.ML Wed Jan 04 14:26:19 2017 +0100
@@ -21,7 +21,7 @@
(* constraints *)
val const_sorts =
- Config.bool (Config.declare ("const_sorts", @{here}) (K (Config.Bool true)));
+ Config.bool (Config.declare ("const_sorts", \<^here>) (K (Config.Bool true)));
fun const_type ctxt =
try ((not (Config.get ctxt const_sorts) ? Type.strip_sorts) o
--- a/src/Pure/unify.ML Wed Jan 04 14:26:08 2017 +0100
+++ b/src/Pure/unify.ML Wed Jan 04 14:26:19 2017 +0100
@@ -32,19 +32,19 @@
(*Unification options*)
(*tracing starts above this depth, 0 for full*)
-val trace_bound_raw = Config.declare ("unify_trace_bound", @{here}) (K (Config.Int 50));
+val trace_bound_raw = Config.declare ("unify_trace_bound", \<^here>) (K (Config.Int 50));
val trace_bound = Config.int trace_bound_raw;
(*unification quits above this depth*)
-val search_bound_raw = Config.declare ("unify_search_bound", @{here}) (K (Config.Int 60));
+val search_bound_raw = Config.declare ("unify_search_bound", \<^here>) (K (Config.Int 60));
val search_bound = Config.int search_bound_raw;
(*print dpairs before calling SIMPL*)
-val trace_simp_raw = Config.declare ("unify_trace_simp", @{here}) (K (Config.Bool false));
+val trace_simp_raw = Config.declare ("unify_trace_simp", \<^here>) (K (Config.Bool false));
val trace_simp = Config.bool trace_simp_raw;
(*announce potential incompleteness of type unification*)
-val trace_types_raw = Config.declare ("unify_trace_types", @{here}) (K (Config.Bool false));
+val trace_types_raw = Config.declare ("unify_trace_types", \<^here>) (K (Config.Bool false));
val trace_types = Config.bool trace_types_raw;
--- a/src/Pure/variable.ML Wed Jan 04 14:26:08 2017 +0100
+++ b/src/Pure/variable.ML Wed Jan 04 14:26:19 2017 +0100
@@ -335,7 +335,7 @@
(* inner body mode *)
val inner_body =
- Config.bool (Config.declare ("inner_body", @{here}) (K (Config.Bool false)));
+ Config.bool (Config.declare ("inner_body", \<^here>) (K (Config.Bool false)));
fun is_body ctxt = Config.get ctxt inner_body;
val set_body = Config.put inner_body;
@@ -345,7 +345,7 @@
(* proper mode *)
val proper_fixes =
- Config.bool (Config.declare ("proper_fixes", @{here}) (K (Config.Bool true)));
+ Config.bool (Config.declare ("proper_fixes", \<^here>) (K (Config.Bool true)));
val improper_fixes = Config.put proper_fixes false;
fun restore_proper_fixes ctxt = Config.put proper_fixes (Config.get ctxt proper_fixes);
@@ -645,7 +645,7 @@
(* focus on outermost parameters: !!x y z. B *)
val bound_focus =
- Config.bool (Config.declare ("bound_focus", @{here}) (K (Config.Bool false)));
+ Config.bool (Config.declare ("bound_focus", \<^here>) (K (Config.Bool false)));
fun is_bound_focus ctxt = Config.get ctxt bound_focus;
val set_bound_focus = Config.put bound_focus;
--- a/src/Sequents/prover.ML Wed Jan 04 14:26:08 2017 +0100
+++ b/src/Sequents/prover.ML Wed Jan 04 14:26:19 2017 +0100
@@ -106,8 +106,8 @@
fun method tac =
Method.sections
- [Args.$$$ "add" -- Args.bang_colon >> K (Method.modifier safe_add @{here}),
- Args.$$$ "add" -- Args.colon >> K (Method.modifier unsafe_add @{here})]
+ [Args.$$$ "add" -- Args.bang_colon >> K (Method.modifier safe_add \<^here>),
+ Args.$$$ "add" -- Args.colon >> K (Method.modifier unsafe_add \<^here>)]
>> K (SIMPLE_METHOD' o tac);
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/Tools/VSCode/README.md Wed Jan 04 14:26:19 2017 +0100
@@ -0,0 +1,63 @@
+# 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))
+
+
+## Run ##
+
+* Extensions: search for "Isabelle", click "Install"
+
+* Preferences / User settings / edit settings.json: e.g.
+ `"isabelle.home": "/home/makarius/isabelle/repos"`
+
+* 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
+
+
+## Build ##
+
+* shell> `cd src/Tools/VSCode/extension`
+
+* shell> `isabelle vscode_grammar`
+
+* shell> `vsce package`
+
+
+## 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/etc/options Wed Jan 04 14:26:19 2017 +0100
@@ -0,0 +1,19 @@
+(* :mode=isabelle-options: *)
+
+option vscode_input_delay : real = 0.3
+ -- "delay for client input (edits)"
+
+option vscode_output_delay : real = 0.5
+ -- "delay for client output (rendering)"
+
+option vscode_load_delay : real = 0.5
+ -- "delay for file load operations"
+
+option vscode_tooltip_margin : int = 60
+ -- "margin for pretty-printing of tooltips"
+
+option vscode_diagnostics_margin : int = 80
+ -- "margin for pretty-printing of diagnostic messages"
+
+option vscode_timing_threshold : real = 0.1
+ -- "default threshold for timing display (seconds)"
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/Tools/VSCode/extension/.vscode/launch.json Wed Jan 04 14:26:19 2017 +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 Wed Jan 04 14:26:19 2017 +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 Wed Jan 04 14:26:19 2017 +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 Wed Jan 04 14:26:19 2017 +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 Wed Jan 04 14:26:19 2017 +0100
@@ -0,0 +1,15 @@
+# Isabelle Prover IDE support
+
+This extension connects to the Isabelle Prover IDE infrastructure, using the
+VSCode Language Server protocol. This requires a recent development version of
+Isabelle from 2017, see also:
+
+ * http://isabelle.in.tum.de/devel
+ * http://isabelle.in.tum.de/repos/isabelle/file/tip/src/Tools/VSCode
+
+
+## Important User Settings ##
+
+ * `isabelle.home` points to the main Isabelle directory (ISABELLE_HOME).
+ * `isabelle.cygwin_root` (on Windows) points to the Cygwin installation,
+ e.g. ISABELLE_HOME/cygwin for a regular Isabelle application bundle.
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/Tools/VSCode/extension/isabelle-language.json Wed Jan 04 14:26:19 2017 +0100
@@ -0,0 +1,48 @@
+{
+ "comments": {
+ "blockComment": ["(*", "*)"]
+ },
+ "brackets": [
+ ["(", ")"],
+ ["[", "]"],
+ ["{", "}"],
+ ["«", "»"],
+ ["‹", "›"],
+ ["⟨", "⟩"],
+ ["⌈", "⌉"],
+ ["⌊", "⌋"],
+ ["⦇", "⦈"],
+ ["⟦", "⟧"],
+ ["⦃", "⦄"]
+ ],
+ "autoClosingPairs": [
+ { "open": "(", "close": ")" },
+ { "open": "[", "close": "]" },
+ { "open": "{", "close": "}" },
+ { "open": "«", "close": "»" },
+ { "open": "‹", "close": "›" },
+ { "open": "⟨", "close": "⟩" },
+ { "open": "⌈", "close": "⌉" },
+ { "open": "⌊", "close": "⌋" },
+ { "open": "⦇", "close": "⦈" },
+ { "open": "⟦", "close": "⟧" },
+ { "open": "⦃", "close": "⦄" },
+ { "open": "`", "close": "`", "notIn": ["string"] },
+ { "open": "\"", "close": "\"", "notIn": ["string"] }
+ ],
+ "surroundingPairs": [
+ ["(", ")"],
+ ["[", "]"],
+ ["{", "}"],
+ ["«", "»"],
+ ["‹", "›"],
+ ["⟨", "⟩"],
+ ["⌈", "⌉"],
+ ["⌊", "⌋"],
+ ["⦇", "⦈"],
+ ["⟦", "⟧"],
+ ["⦃", "⦄"],
+ ["`", "`"],
+ ["\"", "\""]
+ ]
+}
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/Tools/VSCode/extension/isabelle-ml-grammar.json Wed Jan 04 14:26:19 2017 +0100
@@ -0,0 +1,288 @@
+{
+ "name": "Isabelle/ML",
+ "scopeName": "source.isabelle-ml",
+ "fileTypes": ["ML"],
+ "uuid": "aa32eb5e-d0d9-11e6-b7a4-37ba001f1e6e",
+ "keyEquivalent": "^~M",
+ "repository": {
+ "comments": {
+ "patterns": [
+ {
+ "end": "\\*\\)",
+ "begin": "\\(\\*",
+ "beginCaptures": {
+ "0": {
+ "name": "punctuation.definition.comment.begin.ml"
+ }
+ },
+ "patterns": [
+ {
+ "include": "#comments"
+ }
+ ],
+ "endCaptures": {
+ "0": {
+ "name": "punctuation.definition.comment.end.ml"
+ }
+ },
+ "name": "comment.block.ml"
+ }
+ ]
+ },
+ "spec": {
+ "patterns": [
+ {
+ "match": "\\b(exception|type)\\s+([a-zA-Z][a-zA-Z0-9'_]*)",
+ "captures": {
+ "1": {
+ "name": "keyword.other.ml"
+ },
+ "2": {
+ "name": "entity.name.type.abbrev.ml"
+ }
+ },
+ "name": "meta.spec.ml.type"
+ },
+ {
+ "end": "(?=val|type|eqtype|datatype|structure|include|exception)",
+ "begin": "\\b(datatype)\\s+([a-zA-Z][a-zA-Z0-9'_]*)\\s*(?==)",
+ "patterns": [
+ {
+ "match": "\\b(and)\\s+([a-zA-Z][a-zA-Z0-9'_]*)\\s*(?==)",
+ "captures": {
+ "1": {
+ "name": "keyword.other.ml"
+ },
+ "2": {
+ "name": "entity.name.type.datatype.ml"
+ }
+ },
+ "name": "meta.spec.ml.datatype"
+ },
+ {
+ "match": "(?x)\n\t\t\t\t\t\t\t=\\s*([a-zA-Z][a-zA-Z0-9'_]*)(\\s+of)?",
+ "captures": {
+ "1": {
+ "name": "variable.other.dcon.ml"
+ },
+ "2": {
+ "name": "keyword.other.ml"
+ }
+ },
+ "name": "meta.datatype.rule.main.ml"
+ },
+ {
+ "match": "\\|\\s*([a-zA-Z][a-zA-Z0-9'_]*)(\\s+of)?",
+ "captures": {
+ "1": {
+ "name": "variable.other.dcon.ml"
+ },
+ "2": {
+ "name": "keyword.other.ml"
+ }
+ },
+ "name": "meta.datatype.rule.other.ml"
+ }
+ ],
+ "captures": {
+ "1": {
+ "name": "keyword.other.ml"
+ },
+ "2": {
+ "name": "entity.name.type.datatype.ml"
+ }
+ },
+ "name": "meta.spec.ml.datatype"
+ },
+ {
+ "match": "\\b(val)\\s*([^:]+)\\s*:",
+ "captures": {
+ "1": {
+ "name": "keyword.other.ml"
+ }
+ },
+ "name": "meta.spec.ml.val"
+ },
+ {
+ "end": "(?=val|type|eqtype|datatype|structure|include)",
+ "begin": "\\b(structure)\\s*(\\w+)\\s*:",
+ "patterns": [
+ {
+ "match": "\\b(sharing)\\b",
+ "name": "keyword.other.ml"
+ }
+ ],
+ "captures": {
+ "1": {
+ "name": "keyword.other.ml"
+ },
+ "2": {
+ "name": "entity.name.type.module.ml"
+ }
+ },
+ "name": "meta.spec.ml.structure"
+ },
+ {
+ "match": "\\b(include)\\b",
+ "captures": {
+ "1": {
+ "name": "keyword.other.ml"
+ }
+ },
+ "name": "meta.spec.ml.include"
+ },
+ {
+ "include": "#comments"
+ }
+ ]
+ }
+ },
+ "patterns": [
+ {
+ "include": "#comments"
+ },
+ {
+ "match": "\\b(val|datatype|signature|type|op|sharing|struct|as|let|in|abstype|local|where|case|of|fn|raise|exception|handle|ref|infix|infixr|before|end|structure|withtype)\\b",
+ "name": "keyword.other.ml"
+ },
+ {
+ "end": "\\b(end)\\b",
+ "begin": "\\b(let)\\b",
+ "patterns": [
+ {
+ "include": "$self"
+ }
+ ],
+ "captures": {
+ "1": {
+ "name": "keyword.other.ml"
+ },
+ "2": {
+ "name": "keyword.other.ml"
+ }
+ },
+ "name": "meta.exp.let.ml"
+ },
+ {
+ "end": "\\b(end)\\b",
+ "begin": "\\b(sig)\\b",
+ "patterns": [
+ {
+ "include": "#spec"
+ }
+ ],
+ "captures": {
+ "1": {
+ "name": "keyword.other.delimiter.ml"
+ },
+ "2": {
+ "name": "keyword.other.delimiter.ml"
+ }
+ },
+ "name": "meta.module.sigdec.ml"
+ },
+ {
+ "match": "\\b(if|then|else)\\b",
+ "name": "keyword.control.ml"
+ },
+ {
+ "end": "(?=val|type|eqtype|datatype|structure|local)",
+ "begin": "\\b(fun|and)\\s+([\\w]+)\\b",
+ "patterns": [
+ {
+ "include": "source.ml"
+ }
+ ],
+ "captures": {
+ "1": {
+ "name": "keyword.control.fun.ml"
+ },
+ "2": {
+ "name": "entity.name.function.ml"
+ }
+ },
+ "name": "meta.definition.fun.ml"
+ },
+ {
+ "end": "\"",
+ "begin": "\"",
+ "beginCaptures": {
+ "0": {
+ "name": "punctuation.definition.string.begin.ml"
+ }
+ },
+ "patterns": [
+ {
+ "match": "\\\\.",
+ "name": "constant.character.escape.ml"
+ }
+ ],
+ "endCaptures": {
+ "0": {
+ "name": "punctuation.definition.string.end.ml"
+ }
+ },
+ "name": "string.quoted.double.ml"
+ },
+ {
+ "match": "(#\")(\\\\)?.(\")",
+ "captures": {
+ "3": {
+ "name": "punctuation.definition.constant.ml"
+ },
+ "1": {
+ "name": "punctuation.definition.constant.ml"
+ }
+ },
+ "name": "constant.character.ml"
+ },
+ {
+ "match": "\\b\\d*\\.?\\d+\\b",
+ "name": "constant.numeric.ml"
+ },
+ {
+ "match": "\\b(andalso|orelse|not)\\b",
+ "name": "keyword.operator.logical.ml"
+ },
+ {
+ "end": "(?==|:|\\()",
+ "begin": "(?x)\\b\n\t\t\t\t\t(functor|structure|signature)\\s+\n\t\t\t\t\t(\\w+)\\s* # identifier",
+ "captures": {
+ "1": {
+ "name": "storage.type.module.binder.ml"
+ },
+ "2": {
+ "name": "entity.name.type.module.ml"
+ }
+ },
+ "name": "meta.module.dec.ml"
+ },
+ {
+ "match": "\\b(open)\\b",
+ "name": "keyword.other.module.ml"
+ },
+ {
+ "match": "\\b(nil|true|false|NONE|SOME)\\b",
+ "name": "constant.language.ml"
+ },
+ {
+ "end": "$",
+ "begin": "\\b(type|eqtype) .* =",
+ "patterns": [
+ {
+ "match": "(([a-zA-Z0-9\\.\\* ]|(\\->))*)",
+ "name": "meta.typeexp.ml"
+ }
+ ],
+ "captures": {
+ "1": {
+ "name": "keyword.other.typeabbrev.ml"
+ },
+ "2": {
+ "name": "variable.other.typename.ml"
+ }
+ },
+ "name": "meta.typeabbrev.ml"
+ }
+ ]
+}
\ No newline at end of file
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/Tools/VSCode/extension/isabelle-ml-language.json Wed Jan 04 14:26:19 2017 +0100
@@ -0,0 +1,22 @@
+{
+ "comments": {
+ "blockComment": [ "(*", "*)" ]
+ },
+ "brackets": [
+ ["(", ")"],
+ ["[", "]"],
+ ["{", "}"]
+ ],
+ "autoClosingPairs": [
+ { "open": "{", "close": "}" },
+ { "open": "[", "close": "]" },
+ { "open": "(", "close": ")" },
+ { "open": "\"", "close": "\"", "notIn": [ "string" ] }
+ ],
+ "surroundingPairs": [
+ [ "{", "}" ],
+ [ "[", "]" ],
+ [ "(", ")" ],
+ [ "\"", "\"" ]
+ ]
+}
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/package.json Wed Jan 04 14:26:19 2017 +0100
@@ -0,0 +1,89 @@
+{
+ "name": "isabelle",
+ "displayName": "Isabelle",
+ "description": "Isabelle Prover IDE",
+ "keywords": [
+ "theorem prover",
+ "formalized mathematics",
+ "mathematical logic",
+ "functional programming",
+ "document preparation"
+ ],
+ "icon": "isabelle.png",
+ "version": "0.3.0",
+ "publisher": "makarius",
+ "license": "BSD-3-Clause",
+ "repository": { "url": "http://isabelle.in.tum.de/repos/isabelle" },
+ "engines": { "vscode": "^1.8.0" },
+ "categories": ["Languages"],
+ "activationEvents": [
+ "onLanguage:isabelle",
+ "onLanguage:isabelle-ml"
+ ],
+ "main": "./out/src/extension",
+ "contributes": {
+ "languages": [
+ {
+ "id": "isabelle",
+ "aliases": ["Isabelle"],
+ "extensions": [".thy"],
+ "configuration": "./isabelle-language.json"
+ },
+ {
+ "id": "isabelle-ml",
+ "aliases": ["Isabelle/ML"],
+ "extensions": [".ML"],
+ "configuration": "./isabelle-ml-language.json"
+ }
+
+ ],
+ "grammars": [
+ {
+ "language": "isabelle",
+ "scopeName": "source.isabelle",
+ "path": "./isabelle-grammar.json"
+ },
+ {
+ "language": "isabelle-ml",
+ "scopeName": "source.isabelle-ml",
+ "path": "./isabelle-ml-grammar.json"
+ }
+ ],
+ "configuration": {
+ "title": "Isabelle",
+ "properties": {
+ "isabelle.cygwin_root": {
+ "type": "string",
+ "default": "",
+ "description": "Root of Cygwin installation on Windows (e.g. ISABELLE_HOME/cygwin)."
+ },
+ "isabelle.home": {
+ "type": "string",
+ "default": "",
+ "description": "Main Isabelle directory (ISABELLE_HOME)."
+ },
+ "isabelle.args": {
+ "type": "array",
+ "items": { "type": "string" },
+ "default": [],
+ "description": "Command-line arguments for isabelle vscode_server process."
+ }
+ }
+ }
+ },
+ "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 Wed Jan 04 14:26:19 2017 +0100
@@ -0,0 +1,42 @@
+'use strict';
+
+import * as vscode from 'vscode';
+import * as path from 'path';
+import * as os from 'os';
+
+import { LanguageClient, LanguageClientOptions, SettingMonitor, ServerOptions, TransportKind }
+ from 'vscode-languageclient';
+
+
+export function activate(context: vscode.ExtensionContext)
+{
+ let is_windows = os.type().startsWith("Windows")
+
+ let cygwin_root = vscode.workspace.getConfiguration("isabelle").get<string>("cygwin_root");
+ let isabelle_home = vscode.workspace.getConfiguration("isabelle").get<string>("home");
+ let isabelle_args = vscode.workspace.getConfiguration("isabelle").get<Array<string>>("args");
+
+ if (is_windows && cygwin_root == "")
+ vscode.window.showErrorMessage("Missing user settings: isabelle.cygwin_root")
+ else if (isabelle_home == "")
+ vscode.window.showErrorMessage("Missing user settings: isabelle.home")
+ else {
+ let isabelle_tool = isabelle_home.concat("/bin/isabelle")
+ let run =
+ is_windows ?
+ { command: cygwin_root.concat("/bin/bash"),
+ args: ["-l", isabelle_tool, "vscode_server"].concat(isabelle_args) } :
+ { command: isabelle_tool,
+ args: ["vscode_server"].concat(isabelle_args) };
+
+ let server_options: ServerOptions = { run: run, debug: run };
+ let client_options: LanguageClientOptions = {
+ documentSelector: ["isabelle", "isabelle-ml"]
+ };
+
+ let disposable = new LanguageClient("Isabelle", 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 Wed Jan 04 14:26:19 2017 +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 Wed Jan 04 14:26:19 2017 +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 Wed Jan 04 14:26:19 2017 +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 Wed Jan 04 14:26:19 2017 +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 Wed Jan 04 14:26:19 2017 +0100
@@ -0,0 +1,116 @@
+/* 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: Logger = No_Logger)
+{
+ /* 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
+ }
+ }
+
+
+ /* display message */
+
+ def display_message(message_type: Int, msg: String, show: Boolean = true): Unit =
+ write(Protocol.DisplayMessage(message_type, Output.clean_yxml(msg), show))
+
+ def error_message(msg: String) { display_message(Protocol.MessageType.Error, msg, true) }
+ def warning(msg: String) { display_message(Protocol.MessageType.Warning, msg, true) }
+ def writeln(msg: String) { display_message(Protocol.MessageType.Info, msg, true) }
+
+ def log_error_message(msg: String) { display_message(Protocol.MessageType.Error, msg, false) }
+ def log_warning(msg: String) { display_message(Protocol.MessageType.Warning, msg, false) }
+ def log_writeln(msg: String) { display_message(Protocol.MessageType.Info, msg, false) }
+
+
+ /* progress */
+
+ def make_progress(verbose: Boolean = false): Progress =
+ new Progress {
+ override def echo(msg: String): Unit = log_writeln(msg)
+ override def theory(session: String, theory: String): Unit =
+ if (verbose) echo(session + ": theory " + theory)
+ }
+
+
+ /* diagnostics */
+
+ def diagnostics(uri: String, diagnostics: List[Protocol.Diagnostic]): Unit =
+ write(Protocol.PublishDiagnostics(uri, diagnostics))
+}
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/Tools/VSCode/src/document_model.scala Wed Jan 04 14:26:19 2017 +0100
@@ -0,0 +1,139 @@
+/* Title: Tools/VSCode/src/document_model.scala
+ Author: Makarius
+
+Document model for line-oriented text.
+*/
+
+package isabelle.vscode
+
+
+import isabelle._
+
+import java.io.{File => JFile}
+
+import scala.util.parsing.input.CharSequenceReader
+
+
+object Document_Model
+{
+ def init(session: Session, uri: String): Document_Model =
+ {
+ val resources = session.resources.asInstanceOf[VSCode_Resources]
+ val node_name = resources.node_name(uri)
+ val doc = Line.Document("", resources.text_length)
+ Document_Model(session, node_name, doc)
+ }
+}
+
+sealed case class Document_Model(
+ session: Session,
+ node_name: Document.Node.Name,
+ doc: Line.Document,
+ external_file: Boolean = false,
+ node_required: Boolean = false,
+ last_perspective: Document.Node.Perspective_Text = Document.Node.no_perspective_text,
+ pending_edits: Vector[Text.Edit] = Vector.empty,
+ published_diagnostics: List[Text.Info[Command.Results]] = Nil)
+{
+ /* name */
+
+ override def toString: String = node_name.toString
+
+ def uri: String = node_name.node
+ def is_theory: Boolean = node_name.is_theory
+
+
+ /* external file */
+
+ val file: JFile = Url.file(uri).getCanonicalFile
+
+ def external(b: Boolean): Document_Model = copy(external_file = b)
+
+ def register(watcher: File_Watcher)
+ {
+ val dir = file.getParentFile
+ if (dir != null && dir.isDirectory) watcher.register(dir)
+ }
+
+
+ /* header */
+
+ def node_header: Document.Node.Header =
+ resources.special_header(node_name) getOrElse
+ {
+ if (is_theory)
+ resources.check_thy_reader(
+ "", node_name, new CharSequenceReader(Thy_Header.header_text(doc)), Token.Pos.command)
+ else Document.Node.no_header
+ }
+
+
+ /* perspective */
+
+ def node_visible: Boolean = !external_file
+
+ def text_perspective: Text.Perspective =
+ if (node_visible) Text.Perspective.full else Text.Perspective.empty
+
+ def node_perspective: Document.Node.Perspective_Text =
+ Document.Node.Perspective(node_required, text_perspective, Document.Node.Overlays.empty)
+
+
+ /* edits */
+
+ def update_text(new_text: String): Option[Document_Model] =
+ {
+ val old_text = doc.make_text
+ if (new_text == old_text) None
+ else {
+ val doc1 = Line.Document(new_text, doc.text_length)
+ val pending_edits1 =
+ if (old_text != "") pending_edits :+ Text.Edit.remove(0, old_text) else pending_edits
+ val pending_edits2 = pending_edits1 :+ Text.Edit.insert(0, new_text)
+ Some(copy(doc = doc1, pending_edits = pending_edits2))
+ }
+ }
+
+ def update_file: Option[Document_Model] =
+ if (external_file) {
+ try { update_text(File.read(file)) }
+ catch { case ERROR(_) => None }
+ }
+ else None
+
+ def flush_edits: Option[(List[Document.Edit_Text], Document_Model)] =
+ {
+ val perspective = node_perspective
+ if (pending_edits.nonEmpty || last_perspective != perspective) {
+ val text_edits = pending_edits.toList
+ val edits =
+ session.header_edit(node_name, node_header) ::
+ (if (text_edits.isEmpty) Nil
+ else List[Document.Edit_Text](node_name -> Document.Node.Edits(text_edits))) :::
+ (if (last_perspective == perspective) Nil
+ else List[Document.Edit_Text](node_name -> perspective))
+ Some((edits, copy(pending_edits = Vector.empty, last_perspective = perspective)))
+ }
+ else None
+ }
+
+
+ /* diagnostics */
+
+ def publish_diagnostics(rendering: VSCode_Rendering)
+ : Option[(List[Text.Info[Command.Results]], Document_Model)] =
+ {
+ val diagnostics = rendering.diagnostics
+ if (diagnostics == published_diagnostics) None
+ else Some(diagnostics, copy(published_diagnostics = diagnostics))
+ }
+
+
+ /* prover session */
+
+ def resources: VSCode_Resources = session.resources.asInstanceOf[VSCode_Resources]
+
+ def snapshot(): Document.Snapshot = session.snapshot(node_name, pending_edits.toList)
+
+ def rendering(): VSCode_Rendering = new VSCode_Rendering(this, snapshot(), resources)
+}
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/Tools/VSCode/src/grammar.scala Wed Jan 04 14:26:19 2017 +0100
@@ -0,0 +1,171 @@
+/* Title: Tools/VSCode/src/grammar.scala
+ Author: Makarius
+
+Generate static TextMate grammar for VSCode editor.
+*/
+
+package isabelle.vscode
+
+
+import isabelle._
+
+import java.util.UUID
+
+
+object Grammar
+{
+ /* generate grammar */
+
+ private lazy val default_logic = Isabelle_System.getenv("ISABELLE_LOGIC")
+
+ private def default_output(logic: String = ""): String =
+ if (logic == "" || logic == default_logic) "isabelle-grammar.json"
+ else "isabelle-" + logic + "-grammar.json"
+
+ def generate(keywords: Keyword.Keywords): JSON.S =
+ {
+ val (minor_keywords, operators) =
+ keywords.minor.iterator.toList.partition(Symbol.is_ascii_identifier(_))
+
+ def major_keywords(pred: String => Boolean): List[String] =
+ (for {
+ k <- keywords.major.iterator
+ kind <- keywords.kinds.get(k)
+ if pred(kind)
+ } yield k).toList
+
+ val keywords1 =
+ major_keywords(k => k != Keyword.THY_END && k != Keyword.PRF_ASM && k != Keyword.PRF_ASM_GOAL)
+ val keywords2 = minor_keywords ::: major_keywords(Set(Keyword.THY_END))
+ val keywords3 = major_keywords(Set(Keyword.PRF_ASM, Keyword.PRF_ASM_GOAL))
+
+
+ def quote_name(a: String): String =
+ if (Symbol.is_ascii_identifier(a)) a else "\\Q" + a + "\\E"
+
+ def grouped_names(as: List[String]): String =
+ JSON.Format("\\b(" + as.sorted.map(quote_name(_)).mkString("|") + ")\\b")
+
+ """{
+ "name": "Isabelle",
+ "scopeName": "source.isabelle",
+ "fileTypes": ["thy"],
+ "uuid": """ + JSON.Format(UUID.randomUUID().toString) + """,
+ "repository": {
+ "comment": {
+ "patterns": [
+ {
+ "name": "comment.block.isabelle",
+ "begin": "\\(\\*",
+ "patterns": [{ "include": "#comment" }],
+ "end": "\\*\\)"
+ }
+ ]
+ },
+ "cartouche": {
+ "patterns": [
+ {
+ "name": "string.quoted.other.multiline.isabelle",
+ "begin": "(?:\\\\<open>|‹)",
+ "patterns": [{ "include": "#cartouche" }],
+ "end": "(?:\\\\<close>|›)"
+ }
+ ]
+ }
+ },
+ "patterns": [
+ {
+ "include": "#comment"
+ },
+ {
+ "include": "#cartouche"
+ },
+ {
+ "name": "keyword.control.isabelle",
+ "match": """ + grouped_names(keywords1) + """
+ },
+ {
+ "name": "keyword.other.isabelle",
+ "match": """ + grouped_names(keywords2) + """
+ },
+ {
+ "name": "keyword.operator.isabelle",
+ "match": """ + grouped_names(operators) + """
+ },
+ {
+ "name": "constant.language.isabelle",
+ "match": """ + grouped_names(keywords3) + """
+ },
+ {
+ "name": "constant.numeric.isabelle",
+ "match": "\\b\\d*\\.?\\d+\\b"
+ },
+ {
+ "name": "string.quoted.double.isabelle",
+ "begin": "\"",
+ "patterns": [
+ {
+ "name": "constant.character.escape.isabelle",
+ "match": """ + JSON.Format("""\\[\"]|\\\d\d\d""") + """
+ }
+ ],
+ "end": "\""
+ },
+ {
+ "name": "string.quoted.backtick.isabelle",
+ "begin": "`",
+ "patterns": [
+ {
+ "name": "constant.character.escape.isabelle",
+ "match": """ + JSON.Format("""\\[\`]|\\\d\d\d""") + """
+ }
+ ],
+ "end": "`"
+ },
+ {
+ "name": "string.quoted.verbatim.isabelle",
+ "begin": """ + JSON.Format("""\{\*""") + """,
+ "patterns": [
+ { "match": """ + JSON.Format("""[^*]+|\*(?!\})""") + """ }
+ ],
+ "end": """ + JSON.Format("""\*\}""") + """
+ }
+ ]
+}
+"""
+ }
+
+
+ /* Isabelle tool wrapper */
+
+ val isabelle_tool = Isabelle_Tool("vscode_grammar",
+ "generate static TextMate grammar for VSCode editor", args =>
+ {
+ var dirs: List[Path] = Nil
+ var logic = default_logic
+ var output: Option[Path] = None
+
+ val getopts = Getopts("""
+Usage: isabelle vscode_grammar [OPTIONS]
+
+ Options are:
+ -d DIR include session directory
+ -l NAME logic session name (default ISABELLE_LOGIC=""" + quote(default_logic) + """)
+ -o FILE output file name (default """ + default_output() + """)
+
+ Generate static TextMate grammar for VSCode editor.
+""",
+ "d:" -> (arg => dirs = dirs ::: List(Path.explode(arg))),
+ "l:" -> (arg => logic = arg),
+ "o:" -> (arg => output = Some(Path.explode(arg))))
+
+ val more_args = getopts(args)
+ if (more_args.nonEmpty) getopts.usage()
+
+ val keywords = Build.outer_syntax(Options.init(), dirs, logic).keywords
+ val output_path = output getOrElse Path.explode(default_output(logic))
+
+ Output.writeln(output_path.implode)
+ File.write_backup(output_path, generate(keywords))
+ })
+}
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/Tools/VSCode/src/protocol.scala Wed Jan 04 14:26:19 2017 +0100
@@ -0,0 +1,370 @@
+/* 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
+ }
+ }
+
+ class RequestTextDocumentPosition(name: String)
+ {
+ def unapply(json: JSON.T): Option[(Id, Line.Node_Position)] =
+ json match {
+ case RequestMessage(id, method, Some(TextDocumentPosition(node_pos))) if method == name =>
+ Some((id, node_pos))
+ 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,
+ "definitionProvider" -> true,
+ "documentHighlightProvider" -> 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(loc: Line.Node_Range): JSON.T =
+ Map("uri" -> loc.name, "range" -> Range(loc.range))
+
+ def unapply(json: JSON.T): Option[Line.Node_Range] =
+ for {
+ uri <- JSON.string(json, "uri")
+ range_json <- JSON.value(json, "range")
+ range <- Range.unapply(range_json)
+ } yield Line.Node_Range(uri, range)
+ }
+
+ object TextDocumentPosition
+ {
+ def unapply(json: JSON.T): Option[Line.Node_Position] =
+ for {
+ doc <- JSON.value(json, "textDocument")
+ uri <- JSON.string(doc, "uri")
+ pos_json <- JSON.value(json, "position")
+ pos <- Position.unapply(pos_json)
+ } yield Line.Node_Position(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 extends RequestTextDocumentPosition("textDocument/hover")
+ {
+ def reply(id: Id, result: Option[(Line.Range, List[String])]): JSON.T =
+ {
+ val res =
+ result match {
+ case Some((range, contents)) =>
+ Map("contents" -> contents.map(s => Map("language" -> "plaintext", "value" -> s)),
+ "range" -> Range(range))
+ case None => Map("contents" -> Nil)
+ }
+ ResponseMessage(id, Some(res))
+ }
+ }
+
+
+ /* goto definition request */
+
+ object GotoDefinition extends RequestTextDocumentPosition("textDocument/definition")
+ {
+ def reply(id: Id, result: List[Line.Node_Range]): JSON.T =
+ ResponseMessage(id, Some(result.map(Location.apply(_))))
+ }
+
+
+ /* document highlights request */
+
+ object DocumentHighlight
+ {
+ def text(range: Line.Range): DocumentHighlight = DocumentHighlight(range, Some(1))
+ def read(range: Line.Range): DocumentHighlight = DocumentHighlight(range, Some(2))
+ def write(range: Line.Range): DocumentHighlight = DocumentHighlight(range, Some(3))
+ }
+
+ sealed case class DocumentHighlight(range: Line.Range, kind: Option[Int] = None)
+ {
+ def json: JSON.T =
+ kind match {
+ case None => Map("range" -> Range(range))
+ case Some(k) => Map("range" -> Range(range), "kind" -> k)
+ }
+ }
+
+ object DocumentHighlights extends RequestTextDocumentPosition("textDocument/documentHighlight")
+ {
+ def reply(id: Id, result: List[DocumentHighlight]): JSON.T =
+ ResponseMessage(id, Some(result.map(_.json)))
+ }
+
+
+ /* diagnostics */
+
+ sealed case class Diagnostic(range: Line.Range, message: String,
+ severity: Option[Int] = None, code: Option[Int] = None, source: Option[String] = None)
+ {
+ def json: JSON.T =
+ Message.empty + ("range" -> Range(range)) + ("message" -> message) ++
+ (severity match { case Some(x) => Map("severity" -> x) case None => Map.empty }) ++
+ (code match { case Some(x) => Map("code" -> x) case None => Map.empty }) ++
+ (source match { case Some(x) => Map("source" -> x) case None => Map.empty })
+ }
+
+ object DiagnosticSeverity
+ {
+ val Error = 1
+ val Warning = 2
+ val Information = 3
+ val Hint = 4
+ }
+
+ object PublishDiagnostics
+ {
+ def apply(uri: String, diagnostics: List[Diagnostic]): JSON.T =
+ Notification("textDocument/publishDiagnostics",
+ Map("uri" -> uri, "diagnostics" -> diagnostics.map(_.json)))
+ }
+}
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/Tools/VSCode/src/server.scala Wed Jan 04 14:26:19 2017 +0100
@@ -0,0 +1,363 @@
+/* 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, File => JFile}
+
+import scala.annotation.tailrec
+
+
+object Server
+{
+ /* Isabelle tool wrapper */
+
+ private val default_text_length = "UTF-16"
+ private lazy val default_logic = Isabelle_System.getenv("ISABELLE_LOGIC")
+
+ val isabelle_tool = Isabelle_Tool("vscode_server", "VSCode Language Server for PIDE", args =>
+ {
+ try {
+ var log_file: Option[Path] = None
+ var text_length = Text.Length.encoding(default_text_length)
+ var dirs: List[Path] = Nil
+ var logic = default_logic
+ var modes: List[String] = Nil
+ var options = Options.init()
+ var system_mode = false
+
+ def text_length_choice: String =
+ commas(Text.Length.encodings.map(
+ { case (a, _) => if (a == default_text_length) a + " (default)" else a }))
+
+ val getopts = Getopts("""
+Usage: isabelle vscode_server [OPTIONS]
+
+ Options are:
+ -L FILE enable logging on FILE
+ -T LENGTH text length encoding: """ + text_length_choice + """
+ -d DIR include session directory
+ -l NAME logic session name (default ISABELLE_LOGIC=""" + quote(default_logic) + """)
+ -m MODE add print mode for output
+ -o OPTION override Isabelle system OPTION (via NAME=VAL or NAME)
+ -s system build mode for session image
+
+ Run the VSCode Language Server protocol (JSON RPC) over stdin/stdout.
+""",
+ "L:" -> (arg => log_file = Some(Path.explode(File.standard_path(arg)))),
+ "T:" -> (arg => Text.Length.encoding(arg)),
+ "d:" -> (arg => dirs = dirs ::: List(Path.explode(File.standard_path(arg)))),
+ "l:" -> (arg => logic = arg),
+ "m:" -> (arg => modes = arg :: modes),
+ "o:" -> (arg => options = options + arg),
+ "s" -> (_ => system_mode = true))
+
+ val more_args = getopts(args)
+ if (more_args.nonEmpty) getopts.usage()
+
+ val log = Logger.make(log_file)
+ val channel = new Channel(System.in, System.out, log)
+ val server = new Server(channel, options, text_length, logic, dirs, modes, system_mode, log)
+
+ // 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) }
+ }
+ catch {
+ case exn: Throwable =>
+ val channel = new Channel(System.in, System.out, No_Logger)
+ channel.error_message(Exn.message(exn))
+ throw(exn)
+ }
+ })
+}
+
+class Server(
+ channel: Channel,
+ options: Options,
+ text_length: Text.Length = Text.Length.encoding(Server.default_text_length),
+ session_name: String = Server.default_logic,
+ session_dirs: List[Path] = Nil,
+ modes: List[String] = Nil,
+ system_mode: Boolean = false,
+ log: Logger = No_Logger)
+{
+ /* prover session */
+
+ private val session_ = Synchronized(None: Option[Session])
+ def session: Session = session_.value getOrElse error("Server inactive")
+ def resources: VSCode_Resources = session.resources.asInstanceOf[VSCode_Resources]
+
+ def rendering_offset(node_pos: Line.Node_Position): Option[(VSCode_Rendering, Text.Offset)] =
+ for {
+ model <- resources.get_model(node_pos.name)
+ rendering = model.rendering()
+ offset <- model.doc.offset(node_pos.pos)
+ } yield (rendering, offset)
+
+
+ /* input from client or file-system */
+
+ private val delay_input =
+ Standard_Thread.delay_last(options.seconds("vscode_input_delay"))
+ { resources.flush_input(session) }
+
+ private val delay_load =
+ Standard_Thread.delay_last(options.seconds("vscode_load_delay"))
+ { if (resources.resolve_dependencies(session)) delay_input.invoke() }
+
+ private val watcher =
+ File_Watcher(sync_documents(_), options.seconds("vscode_load_delay"))
+
+ private def sync_documents(changed: Set[JFile]): Unit =
+ if (resources.sync_models(changed)) delay_input.invoke()
+
+ private def update_document(uri: String, text: String)
+ {
+ resources.update_model(session, uri, text)
+ delay_input.invoke()
+ }
+
+ private def close_document(uri: String)
+ {
+ resources.close_model(uri) match {
+ case Some(model) =>
+ model.register(watcher)
+ sync_documents(Set(model.file))
+ case None =>
+ }
+ }
+
+
+ /* output to client */
+
+ private val delay_output: Standard_Thread.Delay =
+ Standard_Thread.delay_last(options.seconds("vscode_output_delay"))
+ {
+ if (session.current_state().stable_tip_version.isEmpty) delay_output.invoke()
+ else resources.flush_output(channel)
+ }
+
+ private val prover_output =
+ Session.Consumer[Session.Commands_Changed](getClass.getName) {
+ case changed if changed.nodes.nonEmpty =>
+ resources.update_output(changed.nodes.toList.map(_.node))
+ delay_output.invoke()
+ case _ =>
+ }
+
+ private val syslog =
+ Session.Consumer[Prover.Message](getClass.getName) {
+ case output: Prover.Output if output.is_syslog =>
+ channel.log_writeln(XML.content(output.message))
+ case _ =>
+ }
+
+
+ /* init and exit */
+
+ def init(id: Protocol.Id)
+ {
+ def reply(err: String)
+ {
+ channel.write(Protocol.Initialize.reply(id, err))
+ if (err == "")
+ channel.writeln("Welcome to Isabelle/" + session_name + " (" + Distribution.version + ")")
+ else channel.error_message(err)
+ }
+
+ val try_session =
+ try {
+ if (!Build.build(options = options, build_heap = true, no_build = true,
+ system_mode = system_mode, dirs = session_dirs, sessions = List(session_name)).ok)
+ {
+ val start_msg = "Build started for Isabelle/" + session_name + " ..."
+ val fail_msg = "Session build failed -- prover process remains inactive!"
+
+ val progress = channel.make_progress(verbose = true)
+ progress.echo(start_msg); channel.writeln(start_msg)
+
+ if (!Build.build(options = options, progress = progress, build_heap = true,
+ system_mode = system_mode, dirs = session_dirs, sessions = List(session_name)).ok)
+ {
+ progress.echo(fail_msg); error(fail_msg)
+ }
+ }
+
+ val content = Build.session_content(options, false, session_dirs, session_name)
+ val resources =
+ new VSCode_Resources(options, text_length, content.loaded_theories,
+ content.known_theories, content.syntax, log) {
+ override def commit(change: Session.Change): Unit =
+ if (change.deps_changed) delay_load.invoke()
+ }
+
+ Some(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")
+ })
+ }
+ catch { case ERROR(msg) => reply(msg); None }
+
+ for (session <- try_session) {
+ 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.commands_changed += prover_output
+ session.all_messages += syslog
+
+ session.start(receiver =>
+ Isabelle_Process(options = options, logic = session_name, dirs = session_dirs,
+ modes = modes, receiver = receiver))
+
+ session_.change(_ => Some(session))
+ }
+ }
+
+ def shutdown(id: Protocol.Id)
+ {
+ def reply(err: String): Unit = channel.write(Protocol.Shutdown.reply(id, err))
+
+ session_.change({
+ 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
+ session.commands_changed -= prover_output
+ session.all_messages -= syslog
+ reply("")
+ case _ =>
+ }
+ session.phase_changed += session_phase
+ session.stop()
+ delay_input.revoke()
+ delay_output.revoke()
+ watcher.shutdown()
+ None
+ case None =>
+ reply("Prover inactive")
+ None
+ })
+ }
+
+ def exit() {
+ log("\n")
+ sys.exit(if (session_.value.isDefined) 1 else 0)
+ }
+
+
+ /* hover */
+
+ def hover(id: Protocol.Id, node_pos: Line.Node_Position)
+ {
+ val result =
+ for {
+ (rendering, offset) <- rendering_offset(node_pos)
+ info <- rendering.tooltips(Text.Range(offset, offset + 1))
+ } yield {
+ val doc = rendering.model.doc
+ val range = doc.range(info.range)
+ val contents =
+ info.info.map(tree => Pretty.string_of(List(tree), margin = rendering.tooltip_margin))
+ (range, contents)
+ }
+ channel.write(Protocol.Hover.reply(id, result))
+ }
+
+
+ /* goto definition */
+
+ def goto_definition(id: Protocol.Id, node_pos: Line.Node_Position)
+ {
+ val result =
+ (for ((rendering, offset) <- rendering_offset(node_pos))
+ yield rendering.hyperlinks(Text.Range(offset, offset + 1))) getOrElse Nil
+ channel.write(Protocol.GotoDefinition.reply(id, result))
+ }
+
+
+ /* document highlights */
+
+ def document_highlights(id: Protocol.Id, node_pos: Line.Node_Position)
+ {
+ val result =
+ (for ((rendering, offset) <- rendering_offset(node_pos))
+ yield {
+ val doc = rendering.model.doc
+ rendering.caret_focus_ranges(Text.Range(offset, offset + 1), doc.full_range)
+ .map(r => Protocol.DocumentHighlight.text(doc.range(r)))
+ }) getOrElse Nil
+ channel.write(Protocol.DocumentHighlights.reply(id, result))
+ }
+
+
+ /* main loop */
+
+ def start()
+ {
+ log("Server started " + Date.now())
+
+ def handle(json: JSON.T)
+ {
+ try {
+ json match {
+ case Protocol.Initialize(id) => init(id)
+ case Protocol.Shutdown(id) => shutdown(id)
+ 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) =>
+ close_document(uri)
+ case Protocol.Hover(id, node_pos) => hover(id, node_pos)
+ case Protocol.GotoDefinition(id, node_pos) => goto_definition(id, node_pos)
+ case Protocol.DocumentHighlights(id, node_pos) => document_highlights(id, node_pos)
+ case _ => log("### IGNORED")
+ }
+ }
+ catch { case exn: Throwable => channel.log_error_message(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 => log("### TERMINATE")
+ }
+ }
+ loop()
+ }
+}
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/Tools/VSCode/src/vscode_rendering.scala Wed Jan 04 14:26:19 2017 +0100
@@ -0,0 +1,141 @@
+/* Title: Tools/VSCode/src/vscode_rendering.scala
+ Author: Makarius
+
+Isabelle/VSCode-specific implementation of quasi-abstract rendering and
+markup interpretation.
+*/
+
+package isabelle.vscode
+
+
+import isabelle._
+
+object VSCode_Rendering
+{
+ /* diagnostic messages */
+
+ private val message_severity =
+ Map(
+ Markup.WRITELN -> Protocol.DiagnosticSeverity.Information,
+ Markup.INFORMATION -> Protocol.DiagnosticSeverity.Information,
+ Markup.WARNING -> Protocol.DiagnosticSeverity.Warning,
+ Markup.LEGACY -> Protocol.DiagnosticSeverity.Warning,
+ Markup.ERROR -> Protocol.DiagnosticSeverity.Error,
+ Markup.BAD -> Protocol.DiagnosticSeverity.Error)
+
+
+ /* markup elements */
+
+ private val diagnostics_elements =
+ Markup.Elements(Markup.WRITELN, Markup.INFORMATION, Markup.WARNING, Markup.LEGACY, Markup.ERROR,
+ Markup.BAD)
+
+ private val hyperlink_elements =
+ Markup.Elements(Markup.ENTITY, Markup.PATH, Markup.POSITION)
+}
+
+class VSCode_Rendering(
+ val model: Document_Model,
+ snapshot: Document.Snapshot,
+ resources: VSCode_Resources)
+ extends Rendering(snapshot, resources.options, resources)
+{
+ /* diagnostics */
+
+ def diagnostics: List[Text.Info[Command.Results]] =
+ snapshot.cumulate[Command.Results](
+ model.doc.full_range, Command.Results.empty, VSCode_Rendering.diagnostics_elements, _ =>
+ {
+ case (res, Text.Info(_, msg @ XML.Elem(Markup(_, Markup.Serial(i)), body)))
+ if body.nonEmpty => Some(res + (i -> msg))
+
+ case _ => None
+ })
+
+ val diagnostics_margin = options.int("vscode_diagnostics_margin")
+
+ def diagnostics_output(results: List[Text.Info[Command.Results]]): List[Protocol.Diagnostic] =
+ {
+ (for {
+ Text.Info(text_range, res) <- results.iterator
+ range = model.doc.range(text_range)
+ (_, XML.Elem(Markup(name, _), body)) <- res.iterator
+ } yield {
+ val message = Pretty.string_of(body, margin = diagnostics_margin)
+ val severity = VSCode_Rendering.message_severity.get(name)
+ Protocol.Diagnostic(range, message, severity = severity)
+ }).toList
+ }
+
+
+ /* tooltips */
+
+ def tooltip_margin: Int = options.int("vscode_tooltip_margin")
+ def timing_threshold: Double = options.real("vscode_timing_threshold")
+
+
+ /* hyperlinks */
+
+ def hyperlink_source_file(source_name: String, line1: Int, range: Symbol.Range)
+ : Option[Line.Node_Range] =
+ {
+ for (name <- resources.source_file(source_name))
+ yield {
+ val opt_text =
+ try { Some(File.read(Path.explode(name))) } // FIXME content from resources/models
+ catch { case ERROR(_) => None }
+ Line.Node_Range(Url.platform_file(name),
+ opt_text match {
+ case Some(text) if range.start > 0 =>
+ val chunk = Symbol.Text_Chunk(text)
+ val doc = Line.Document(text, resources.text_length)
+ doc.range(chunk.decode(range))
+ case _ =>
+ Line.Range(Line.Position((line1 - 1) max 0))
+ })
+ }
+ }
+
+ def hyperlink_command(id: Document_ID.Generic, range: Symbol.Range): Option[Line.Node_Range] =
+ {
+ if (snapshot.is_outdated) None
+ else
+ for {
+ start <- snapshot.find_command_position(id, range.start)
+ stop <- snapshot.find_command_position(id, range.stop)
+ } yield Line.Node_Range(start.name, Line.Range(start.pos, stop.pos))
+ }
+
+ def hyperlink_position(pos: Position.T): Option[Line.Node_Range] =
+ pos match {
+ case Position.Item_File(name, line, range) => hyperlink_source_file(name, line, range)
+ case Position.Item_Id(id, range) => hyperlink_command(id, range)
+ case _ => None
+ }
+
+ def hyperlink_def_position(pos: Position.T): Option[Line.Node_Range] =
+ pos match {
+ case Position.Item_Def_File(name, line, range) => hyperlink_source_file(name, line, range)
+ case Position.Item_Def_Id(id, range) => hyperlink_command(id, range)
+ case _ => None
+ }
+
+ def hyperlinks(range: Text.Range): List[Line.Node_Range] =
+ {
+ snapshot.cumulate[List[Line.Node_Range]](
+ range, Nil, VSCode_Rendering.hyperlink_elements, _ =>
+ {
+ case (links, Text.Info(_, XML.Elem(Markup.Path(name), _))) =>
+ val file = resources.append_file(snapshot.node_name.master_dir, name)
+ Some(Line.Node_Range(file) :: links)
+
+ case (links, Text.Info(info_range, XML.Elem(Markup(Markup.ENTITY, props), _))) =>
+ hyperlink_def_position(props).map(_ :: links)
+
+ case (links, Text.Info(info_range, XML.Elem(Markup(Markup.POSITION, props), _))) =>
+ hyperlink_position(props).map(_ :: links)
+
+ case _ => None
+ }) match { case Text.Info(_, links) :: _ => links.reverse case _ => Nil }
+ }
+}
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/Tools/VSCode/src/vscode_resources.scala Wed Jan 04 14:26:19 2017 +0100
@@ -0,0 +1,201 @@
+/* Title: Tools/VSCode/src/vscode_resources.scala
+ Author: Makarius
+
+Resources for VSCode Language Server: file-system access and global state.
+*/
+
+package isabelle.vscode
+
+
+import isabelle._
+
+import java.io.{File => JFile}
+
+import scala.util.parsing.input.{Reader, CharSequenceReader}
+
+
+object VSCode_Resources
+{
+ /* internal state */
+
+ sealed case class State(
+ models: Map[String, Document_Model] = Map.empty,
+ pending_input: Set[String] = Set.empty,
+ pending_output: Set[String] = Set.empty)
+}
+
+class VSCode_Resources(
+ val options: Options,
+ val text_length: Text.Length,
+ loaded_theories: Set[String],
+ known_theories: Map[String, Document.Node.Name],
+ base_syntax: Outer_Syntax,
+ log: Logger = No_Logger)
+ extends Resources(loaded_theories, known_theories, base_syntax, log)
+{
+ private val state = Synchronized(VSCode_Resources.State())
+
+
+ /* document node name */
+
+ def node_name(uri: String): Document.Node.Name =
+ {
+ val theory = Thy_Header.thy_name_bootstrap(uri).getOrElse("")
+ val master_dir =
+ if (!Url.is_wellformed_file(uri) || theory == "") ""
+ else Thy_Header.dir_name(uri)
+ Document.Node.Name(uri, master_dir, theory)
+ }
+
+ override def append(dir: String, source_path: Path): String =
+ {
+ val path = source_path.expand
+ if (path.is_absolute) Url.platform_file(path)
+ else if (dir == "") Url.platform_file(File.pwd() + path)
+ else if (path.is_current) dir
+ else Url.normalize_file(dir + "/" + path.implode)
+ }
+
+ override def with_thy_reader[A](name: Document.Node.Name, f: Reader[Char] => A): A =
+ {
+ val uri = name.node
+ get_model(uri) match {
+ case Some(model) =>
+ val reader = new CharSequenceReader(model.doc.make_text)
+ f(reader)
+
+ case None =>
+ val file = Url.file(uri)
+ if (!file.isFile) error("No such file: " + quote(file.toString))
+
+ val reader = Scan.byte_reader(file)
+ try { f(reader) } finally { reader.close }
+ }
+ }
+
+
+ /* document models */
+
+ def get_model(uri: String): Option[Document_Model] = state.value.models.get(uri)
+
+ def update_model(session: Session, uri: String, text: String)
+ {
+ state.change(st =>
+ {
+ val model = st.models.getOrElse(uri, Document_Model.init(session, uri))
+ val model1 = (model.update_text(text) getOrElse model).external(false)
+ st.copy(
+ models = st.models + (uri -> model1),
+ pending_input = st.pending_input + uri)
+ })
+ }
+
+ def close_model(uri: String): Option[Document_Model] =
+ state.change_result(st =>
+ st.models.get(uri) match {
+ case None => (None, st)
+ case Some(model) =>
+ (Some(model), st.copy(models = st.models + (uri -> model.external(true))))
+ })
+
+ def sync_models(changed_files: Set[JFile]): Boolean =
+ state.change_result(st =>
+ {
+ val changed_models =
+ (for {
+ (uri, model) <- st.models.iterator
+ if changed_files(model.file)
+ model1 <- model.update_file
+ } yield (uri, model1)).toList
+ if (changed_models.isEmpty) (false, st)
+ else (true,
+ st.copy(
+ models = (st.models /: changed_models)(_ + _),
+ pending_input = (st.pending_input /: changed_models.iterator.map(_._1))(_ + _)))
+ })
+
+
+ /* resolve dependencies */
+
+ val thy_info = new Thy_Info(this)
+
+ def resolve_dependencies(session: Session): Boolean =
+ {
+ state.change_result(st =>
+ {
+ val thys =
+ (for ((_, model) <- st.models.iterator if model.is_theory)
+ yield (model.node_name, Position.none)).toList
+
+ val loaded_models =
+ (for {
+ dep <- thy_info.dependencies("", thys).deps.iterator
+ uri = dep.name.node
+ if !st.models.isDefinedAt(uri)
+ text <-
+ try { Some(File.read(Url.file(uri))) }
+ catch { case ERROR(_) => None }
+ }
+ yield {
+ val model = Document_Model.init(session, uri)
+ val model1 = (model.update_text(text) getOrElse model).external(true)
+ (uri, model1)
+ }).toList
+
+ if (loaded_models.isEmpty) (false, st)
+ else
+ (true,
+ st.copy(
+ models = st.models ++ loaded_models,
+ pending_input = st.pending_input ++ loaded_models.iterator.map(_._1)))
+ })
+ }
+
+
+ /* pending input */
+
+ def flush_input(session: Session)
+ {
+ state.change(st =>
+ {
+ val changed_models =
+ (for {
+ uri <- st.pending_input.iterator
+ model <- st.models.get(uri)
+ res <- model.flush_edits
+ } yield res).toList
+
+ session.update(Document.Blobs.empty, changed_models.flatMap(_._1))
+ st.copy(
+ models = (st.models /: changed_models) { case (ms, (_, m)) => ms + (m.uri -> m) },
+ pending_input = Set.empty)
+ })
+ }
+
+
+ /* pending output */
+
+ def update_output(changed_nodes: List[String]): Unit =
+ state.change(st => st.copy(pending_output = st.pending_output ++ changed_nodes))
+
+ def flush_output(channel: Channel)
+ {
+ state.change(st =>
+ {
+ val changed_iterator =
+ for {
+ uri <- st.pending_output.iterator
+ model <- st.models.get(uri)
+ rendering = model.rendering()
+ (diagnostics, model1) <- model.publish_diagnostics(rendering)
+ } yield {
+ channel.diagnostics(model1.uri, rendering.diagnostics_output(diagnostics))
+ model1
+ }
+ st.copy(
+ models = (st.models /: changed_iterator) { case (ms, m) => ms + (m.uri -> m) },
+ pending_output = Set.empty)
+ }
+ )
+ }
+}
--- a/src/Tools/intuitionistic.ML Wed Jan 04 14:26:08 2017 +0100
+++ b/src/Tools/intuitionistic.ML Wed Jan 04 14:26:19 2017 +0100
@@ -78,13 +78,13 @@
>> (fn arg => Method.modifier (att arg) pos);
val modifiers =
- [modifier destN Args.bang_colon Args.bang Context_Rules.dest_bang @{here},
- modifier destN Args.colon (Scan.succeed ()) Context_Rules.dest @{here},
- modifier elimN Args.bang_colon Args.bang Context_Rules.elim_bang @{here},
- modifier elimN Args.colon (Scan.succeed ()) Context_Rules.elim @{here},
- modifier introN Args.bang_colon Args.bang Context_Rules.intro_bang @{here},
- modifier introN Args.colon (Scan.succeed ()) Context_Rules.intro @{here},
- Args.del -- Args.colon >> K (Method.modifier Context_Rules.rule_del @{here})];
+ [modifier destN Args.bang_colon Args.bang Context_Rules.dest_bang \<^here>,
+ modifier destN Args.colon (Scan.succeed ()) Context_Rules.dest \<^here>,
+ modifier elimN Args.bang_colon Args.bang Context_Rules.elim_bang \<^here>,
+ modifier elimN Args.colon (Scan.succeed ()) Context_Rules.elim \<^here>,
+ modifier introN Args.bang_colon Args.bang Context_Rules.intro_bang \<^here>,
+ modifier introN Args.colon (Scan.succeed ()) Context_Rules.intro \<^here>,
+ Args.del -- Args.colon >> K (Method.modifier Context_Rules.rule_del \<^here>)];
in
--- a/src/Tools/jEdit/lib/Tools/jedit Wed Jan 04 14:26:08 2017 +0100
+++ b/src/Tools/jEdit/lib/Tools/jedit Wed Jan 04 14:26:19 2017 +0100
@@ -40,6 +40,7 @@
"src/jedit_editor.scala"
"src/jedit_lib.scala"
"src/jedit_options.scala"
+ "src/jedit_rendering.scala"
"src/jedit_resources.scala"
"src/jedit_sessions.scala"
"src/keymap_merge.scala"
@@ -53,7 +54,6 @@
"src/protocol_dockable.scala"
"src/query_dockable.scala"
"src/raw_output_dockable.scala"
- "src/rendering.scala"
"src/rich_text_area.scala"
"src/scala_console.scala"
"src/session_build.scala"
@@ -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/active.scala Wed Jan 04 14:26:08 2017 +0100
+++ b/src/Tools/jEdit/src/active.scala Wed Jan 04 14:26:19 2017 +0100
@@ -51,7 +51,7 @@
case XML.Elem(Markup(Markup.SIMP_TRACE_PANEL, props), _) =>
val link =
props match {
- case Position.Id(id) => PIDE.editor.hyperlink_command_id(true, snapshot, id)
+ case Position.Id(id) => PIDE.editor.hyperlink_command(true, snapshot, id)
case _ => None
}
GUI_Thread.later {
--- a/src/Tools/jEdit/src/bibtex_jedit.scala Wed Jan 04 14:26:08 2017 +0100
+++ b/src/Tools/jEdit/src/bibtex_jedit.scala Wed Jan 04 14:26:19 2017 +0100
@@ -75,7 +75,7 @@
def completion(
history: Completion.History,
text_area: JEditTextArea,
- rendering: Rendering): Option[Completion.Result] =
+ rendering: JEdit_Rendering): Option[Completion.Result] =
{
for {
Text.Info(r, name) <- rendering.citation(JEdit_Lib.before_caret_range(text_area, rendering))
--- a/src/Tools/jEdit/src/completion_popup.scala Wed Jan 04 14:26:08 2017 +0100
+++ b/src/Tools/jEdit/src/completion_popup.scala Wed Jan 04 14:26:19 2017 +0100
@@ -120,7 +120,7 @@
/* rendering */
- def rendering(rendering: Rendering, line_range: Text.Range): Option[Text.Info[Color]] =
+ def rendering(rendering: JEdit_Rendering, line_range: Text.Range): Option[Text.Info[Color]] =
{
active_range match {
case Some(range) => range.try_restrict(line_range)
@@ -153,7 +153,7 @@
def syntax_completion(
history: Completion.History,
explicit: Boolean,
- opt_rendering: Option[Rendering]): Option[Completion.Result] =
+ opt_rendering: Option[JEdit_Rendering]): Option[Completion.Result] =
{
val buffer = text_area.getBuffer
val decode = Isabelle_Encoding.is_active(buffer)
@@ -185,7 +185,7 @@
/* path completion */
- def path_completion(rendering: Rendering): Option[Completion.Result] =
+ def path_completion(rendering: JEdit_Rendering): Option[Completion.Result] =
{
def complete(text: String): List[(String, List[String])] =
{
@@ -750,7 +750,7 @@
val size =
{
val geometry = JEdit_Lib.window_geometry(completion, completion)
- val bounds = Rendering.popup_bounds
+ val bounds = JEdit_Rendering.popup_bounds
val w = geometry.width min (screen.bounds.width * bounds).toInt min layered.getWidth
val h = geometry.height min (screen.bounds.height * bounds).toInt min layered.getHeight
new Dimension(w, h)
--- a/src/Tools/jEdit/src/document_model.scala Wed Jan 04 14:26:08 2017 +0100
+++ b/src/Tools/jEdit/src/document_model.scala Wed Jan 04 14:26:19 2017 +0100
@@ -69,29 +69,20 @@
class Document_Model(val session: Session, val buffer: Buffer, val node_name: Document.Node.Name)
{
+ override def toString: String = node_name.toString
+
+
/* header */
def is_theory: Boolean = node_name.is_theory
- def is_ml_root: Boolean =
- Thy_Header.ml_roots.exists({ case (_, name) => name == node_name.theory })
-
- def is_bootstrap_thy: Boolean =
- Thy_Header.bootstrap_thys.exists({ case (_, name) => name == node_name.theory })
-
def node_header(): Document.Node.Header =
{
GUI_Thread.require {}
- if (is_ml_root)
- Document.Node.Header(
- List((PIDE.resources.import_name("", node_name, Thy_Header.ML_BOOTSTRAP), Position.none)))
- else if (is_theory) {
- if (is_bootstrap_thy) {
- Document.Node.Header(
- List((PIDE.resources.import_name("", node_name, Thy_Header.PURE), Position.none)))
- }
- else {
+ PIDE.resources.special_header(node_name) getOrElse
+ {
+ if (is_theory) {
JEdit_Lib.buffer_lock(buffer) {
Token_Markup.line_token_iterator(
Thy_Header.bootstrap_syntax, buffer, 0, buffer.getLineCount).collectFirst(
@@ -103,12 +94,13 @@
val length = buffer.getLength - offset
PIDE.resources.check_thy_reader("", node_name,
new CharSequenceReader(buffer.getSegment(offset, length)), Token.Pos.command)
- case None => Document.Node.no_header
+ case None =>
+ Document.Node.no_header
}
}
}
+ else Document.Node.no_header
}
- else Document.Node.no_header
}
--- a/src/Tools/jEdit/src/document_view.scala Wed Jan 04 14:26:08 2017 +0100
+++ b/src/Tools/jEdit/src/document_view.scala Wed Jan 04 14:26:19 2017 +0100
@@ -60,7 +60,7 @@
{
private val session = model.session
- def get_rendering(): Rendering = Rendering(model.snapshot(), PIDE.options.value)
+ def get_rendering(): JEdit_Rendering = JEdit_Rendering(model.snapshot(), PIDE.options.value)
val rich_text_area =
new Rich_Text_Area(text_area.getView, text_area, get_rendering _, () => (), () => None,
--- a/src/Tools/jEdit/src/fold_handling.scala Wed Jan 04 14:26:08 2017 +0100
+++ b/src/Tools/jEdit/src/fold_handling.scala Wed Jan 04 14:26:19 2017 +0100
@@ -50,7 +50,7 @@
/* output: static document rendering */
- class Document_Fold_Handler(private val rendering: Rendering)
+ class Document_Fold_Handler(private val rendering: JEdit_Rendering)
extends FoldHandler("isabelle-document")
{
override def equals(other: Any): Boolean =
--- a/src/Tools/jEdit/src/graphview_dockable.scala Wed Jan 04 14:26:08 2017 +0100
+++ b/src/Tools/jEdit/src/graphview_dockable.scala Wed Jan 04 14:26:19 2017 +0100
@@ -68,7 +68,7 @@
{
Pretty_Tooltip.invoke(() =>
{
- val rendering = Rendering(snapshot, options)
+ val rendering = JEdit_Rendering(snapshot, options)
val info = Text.Info(Text.Range(~1), body)
Pretty_Tooltip(view, parent, new Point(x, y), rendering, Command.Results.empty, info)
})
--- a/src/Tools/jEdit/src/isabelle.scala Wed Jan 04 14:26:08 2017 +0100
+++ b/src/Tools/jEdit/src/isabelle.scala Wed Jan 04 14:26:19 2017 +0100
@@ -329,7 +329,7 @@
{
val buffer = text_area.getBuffer
if (!snapshot.is_outdated && text != "") {
- (snapshot.state.find_command(snapshot.version, id), PIDE.document_model(buffer)) match {
+ (snapshot.find_command(id), PIDE.document_model(buffer)) match {
case (Some((node, command)), Some(model)) if command.node_name == model.node_name =>
node.command_start(command) match {
case Some(start) =>
--- a/src/Tools/jEdit/src/jedit_editor.scala Wed Jan 04 14:26:08 2017 +0100
+++ b/src/Tools/jEdit/src/jedit_editor.scala Wed Jan 04 14:26:19 2017 +0100
@@ -158,21 +158,28 @@
}
}
- def goto_file(focus: Boolean, view: View, name: String, line: Int = 0, column: Int = 0)
+ def goto_file(focus: Boolean, view: View, name: String): Unit =
+ goto_file(focus, view, Line.Node_Position(name))
+
+ def goto_file(focus: Boolean, view: View, pos: Line.Node_Position)
{
GUI_Thread.require {}
push_position(view)
+ val name = pos.name
+ val line = pos.line
+ val column = pos.column
+
JEdit_Lib.jedit_buffer(name) match {
case Some(buffer) =>
if (focus) view.goToBuffer(buffer) else view.showBuffer(buffer)
val text_area = view.getTextArea
try {
- val line_start = text_area.getBuffer.getLineStartOffset(line - 1)
+ val line_start = text_area.getBuffer.getLineStartOffset(line)
text_area.moveCaretPosition(line_start)
- if (column > 0) text_area.moveCaretPosition(line_start + column - 1)
+ if (column > 0) text_area.moveCaretPosition(line_start + column)
}
catch {
case _: ArrayIndexOutOfBoundsException =>
@@ -185,8 +192,8 @@
case None =>
val args =
if (line <= 0) Array(name)
- else if (column <= 0) Array(name, "+line:" + line.toInt)
- else Array(name, "+line:" + line.toInt + "," + column.toInt)
+ else if (column <= 0) Array(name, "+line:" + (line + 1))
+ else Array(name, "+line:" + (line + 1) + "," + (column + 1))
jEdit.openFiles(view, null, args)
}
}
@@ -217,14 +224,14 @@
case doc: Doc.Text_File if doc.name == name => doc.path
case doc: Doc.Doc if doc.name == name => doc.path}).map(path =>
new Hyperlink {
- val external = !path.is_file
+ override val external = !path.is_file
def follow(view: View): Unit = goto_doc(view, path)
override def toString: String = "doc " + quote(name)
})
def hyperlink_url(name: String): Hyperlink =
new Hyperlink {
- val external = true
+ override val external = true
def follow(view: View): Unit =
Standard_Thread.fork("hyperlink_url") {
try { Isabelle_System.open(Url.escape(name)) }
@@ -240,88 +247,56 @@
def hyperlink_buffer(focus: Boolean, buffer: Buffer, offset: Text.Offset): Hyperlink =
new Hyperlink {
- val external = false
def follow(view: View): Unit = goto_buffer(focus, view, buffer, offset)
override def toString: String = "buffer " + quote(JEdit_Lib.buffer_name(buffer))
}
- def hyperlink_file(focus: Boolean, name: String, line: Int = 0, column: Int = 0): Hyperlink =
+ def hyperlink_file(focus: Boolean, name: String): Hyperlink =
+ hyperlink_file(focus, Line.Node_Position(name))
+
+ def hyperlink_file(focus: Boolean, pos: Line.Node_Position): Hyperlink =
new Hyperlink {
- val external = false
- def follow(view: View): Unit = goto_file(focus, view, name, line, column)
- override def toString: String = "file " + quote(name)
+ def follow(view: View): Unit = goto_file(focus, view, pos)
+ override def toString: String = "file " + quote(pos.name)
}
- def hyperlink_source_file(focus: Boolean, source_name: String, line: Int, offset: Symbol.Offset)
+ def hyperlink_source_file(focus: Boolean, source_name: String, line1: Int, offset: Symbol.Offset)
: Option[Hyperlink] =
{
- val opt_name =
- if (Path.is_wellformed(source_name)) {
- if (Path.is_valid(source_name)) {
- val path = Path.explode(source_name)
- Some(File.platform_path(Isabelle_System.source_file(path) getOrElse path))
- }
- else None
+ for (name <- PIDE.resources.source_file(source_name)) yield {
+ JEdit_Lib.jedit_buffer(name) match {
+ case Some(buffer) if offset > 0 =>
+ val pos =
+ JEdit_Lib.buffer_lock(buffer) {
+ (Line.Position.zero /:
+ (Symbol.iterator(JEdit_Lib.buffer_text(buffer)).
+ zipWithIndex.takeWhile(p => p._2 < offset - 1).
+ map(_._1)))(_.advance(_, Text.Length))
+ }
+ hyperlink_file(focus, Line.Node_Position(name, pos))
+ case _ =>
+ hyperlink_file(focus, Line.Node_Position(name, Line.Position((line1 - 1) max 0)))
}
- else Some(source_name)
-
- opt_name match {
- case Some(name) =>
- JEdit_Lib.jedit_buffer(name) match {
- case Some(buffer) if offset > 0 =>
- val (line, column) =
- JEdit_Lib.buffer_lock(buffer) {
- ((1, 1) /:
- (Symbol.iterator(JEdit_Lib.buffer_text(buffer)).
- zipWithIndex.takeWhile(p => p._2 < offset - 1).map(_._1)))(
- Symbol.advance_line_column)
- }
- Some(hyperlink_file(focus, name, line, column))
- case _ => Some(hyperlink_file(focus, name, line))
- }
- case None => None
}
}
override def hyperlink_command(
- focus: Boolean, snapshot: Document.Snapshot, command: Command, offset: Symbol.Offset = 0)
+ focus: Boolean, snapshot: Document.Snapshot, id: Document_ID.Generic, offset: Symbol.Offset = 0)
: Option[Hyperlink] =
{
if (snapshot.is_outdated) None
- else {
- snapshot.state.find_command(snapshot.version, command.id) match {
- case None => None
- case Some((node, _)) =>
- val file_name = command.node_name.node
- val sources_iterator =
- node.commands.iterator.takeWhile(_ != command).map(_.source) ++
- (if (offset == 0) Iterator.empty
- else Iterator.single(command.source(Text.Range(0, command.chunk.decode(offset)))))
- val (line, column) = ((1, 1) /: sources_iterator)(Symbol.advance_line_column)
- Some(hyperlink_file(focus, file_name, line, column))
- }
- }
- }
-
- def hyperlink_command_id(
- focus: Boolean, snapshot: Document.Snapshot, id: Document_ID.Generic, offset: Symbol.Offset = 0)
- : Option[Hyperlink] =
- {
- snapshot.state.find_command(snapshot.version, id) match {
- case Some((node, command)) => hyperlink_command(focus, snapshot, command, offset)
- case None => None
- }
+ else snapshot.find_command_position(id, offset).map(hyperlink_file(focus, _))
}
def is_hyperlink_position(snapshot: Document.Snapshot,
text_offset: Text.Offset, pos: Position.T): Boolean =
{
pos match {
- case Position.Id_Offset0(id, offset) if offset > 0 =>
- snapshot.state.find_command(snapshot.version, id) match {
+ case Position.Item_Id(id, range) if range.start > 0 =>
+ snapshot.find_command(id) match {
case Some((node, command)) if snapshot.version.nodes(command.node_name) eq node =>
node.command_start(command) match {
- case Some(start) => text_offset == start + command.chunk.decode(offset)
+ case Some(start) => text_offset == start + command.chunk.decode(range.start)
case None => false
}
case _ => false
@@ -333,22 +308,20 @@
def hyperlink_position(focus: Boolean, snapshot: Document.Snapshot, pos: Position.T)
: Option[Hyperlink] =
pos match {
- case Position.Line_File(line, name) =>
- val offset = Position.Offset.unapply(pos) getOrElse 0
- hyperlink_source_file(focus, name, line, offset)
- case Position.Id_Offset0(id, offset) =>
- hyperlink_command_id(focus, snapshot, id, offset)
+ case Position.Item_File(name, line, range) =>
+ hyperlink_source_file(focus, name, line, range.start)
+ case Position.Item_Id(id, range) =>
+ hyperlink_command(focus, snapshot, id, range.start)
case _ => None
}
def hyperlink_def_position(focus: Boolean, snapshot: Document.Snapshot, pos: Position.T)
: Option[Hyperlink] =
pos match {
- case Position.Def_Line_File(line, name) =>
- val offset = Position.Def_Offset.unapply(pos) getOrElse 0
- hyperlink_source_file(focus, name, line, offset)
- case Position.Def_Id_Offset0(id, offset) =>
- hyperlink_command_id(focus, snapshot, id, offset)
+ case Position.Item_Def_File(name, line, range) =>
+ hyperlink_source_file(focus, name, line, range.start)
+ case Position.Item_Def_Id(id, range) =>
+ hyperlink_command(focus, snapshot, id, range.start)
case _ => None
}
}
--- a/src/Tools/jEdit/src/jedit_lib.scala Wed Jan 04 14:26:08 2017 +0100
+++ b/src/Tools/jEdit/src/jedit_lib.scala Wed Jan 04 14:26:19 2017 +0100
@@ -198,7 +198,7 @@
def caret_range(text_area: TextArea): Text.Range =
point_range(text_area.getBuffer, text_area.getCaretPosition)
- def before_caret_range(text_area: TextArea, rendering: Rendering): Text.Range =
+ def before_caret_range(text_area: TextArea, rendering: JEdit_Rendering): Text.Range =
{
val snapshot = rendering.snapshot
val former_caret = snapshot.revert(text_area.getCaretPosition)
@@ -320,7 +320,7 @@
{
val name1 =
if (name.startsWith("idea-icons/")) {
- val file = File.platform_url(Path.explode("$JEDIT_HOME/dist/jars/idea-icons.jar"))
+ val file = Path.explode("$JEDIT_HOME/dist/jars/idea-icons.jar").file.toURI.toASCIIString
"jar:" + file + "!/" + name
}
else name
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/Tools/jEdit/src/jedit_rendering.scala Wed Jan 04 14:26:19 2017 +0100
@@ -0,0 +1,724 @@
+/* Title: Tools/jEdit/src/jedit_rendering.scala
+ Author: Makarius
+
+Isabelle/jEdit-specific implementation of quasi-abstract rendering and
+markup interpretation.
+*/
+
+package isabelle.jedit
+
+
+import isabelle._
+
+import java.awt.Color
+import javax.swing.Icon
+
+import org.gjt.sp.jedit.syntax.{Token => JEditToken}
+import org.gjt.sp.jedit.jEdit
+
+import scala.collection.immutable.SortedMap
+
+
+object JEdit_Rendering
+{
+ def apply(snapshot: Document.Snapshot, options: Options): JEdit_Rendering =
+ new JEdit_Rendering(snapshot, options)
+
+
+ /* popup window bounds */
+
+ def popup_bounds: Double = (PIDE.options.real("jedit_popup_bounds") max 0.2) min 0.8
+
+
+ /* Isabelle/Isar token markup */
+
+ private val command_style: Map[String, Byte] =
+ {
+ import JEditToken._
+ Map[String, Byte](
+ Keyword.THY_END -> KEYWORD2,
+ Keyword.PRF_ASM -> KEYWORD3,
+ Keyword.PRF_ASM_GOAL -> KEYWORD3
+ ).withDefaultValue(KEYWORD1)
+ }
+
+ private val token_style: Map[Token.Kind.Value, Byte] =
+ {
+ import JEditToken._
+ Map[Token.Kind.Value, Byte](
+ Token.Kind.KEYWORD -> KEYWORD2,
+ Token.Kind.IDENT -> NULL,
+ Token.Kind.LONG_IDENT -> NULL,
+ Token.Kind.SYM_IDENT -> NULL,
+ Token.Kind.VAR -> NULL,
+ Token.Kind.TYPE_IDENT -> NULL,
+ Token.Kind.TYPE_VAR -> NULL,
+ Token.Kind.NAT -> NULL,
+ Token.Kind.FLOAT -> NULL,
+ Token.Kind.SPACE -> NULL,
+ Token.Kind.STRING -> LITERAL1,
+ Token.Kind.ALT_STRING -> LITERAL2,
+ Token.Kind.VERBATIM -> COMMENT3,
+ Token.Kind.CARTOUCHE -> COMMENT4,
+ Token.Kind.COMMENT -> COMMENT1,
+ Token.Kind.ERROR -> INVALID
+ ).withDefaultValue(NULL)
+ }
+
+ def token_markup(syntax: Outer_Syntax, token: Token): Byte =
+ if (token.is_command) command_style(syntax.keywords.kinds.getOrElse(token.content, ""))
+ else if (token.is_keyword && token.source == Symbol.comment_decoded) JEditToken.NULL
+ else if (token.is_delimiter) JEditToken.OPERATOR
+ else token_style(token.kind)
+
+
+ /* Isabelle/ML token markup */
+
+ private val ml_token_style: Map[ML_Lex.Kind.Value, Byte] =
+ {
+ import JEditToken._
+ Map[ML_Lex.Kind.Value, Byte](
+ ML_Lex.Kind.KEYWORD -> NULL,
+ ML_Lex.Kind.IDENT -> NULL,
+ ML_Lex.Kind.LONG_IDENT -> NULL,
+ ML_Lex.Kind.TYPE_VAR -> NULL,
+ ML_Lex.Kind.WORD -> DIGIT,
+ ML_Lex.Kind.INT -> DIGIT,
+ ML_Lex.Kind.REAL -> DIGIT,
+ ML_Lex.Kind.CHAR -> LITERAL2,
+ ML_Lex.Kind.STRING -> LITERAL1,
+ ML_Lex.Kind.SPACE -> NULL,
+ ML_Lex.Kind.COMMENT -> COMMENT1,
+ ML_Lex.Kind.ANTIQ -> NULL,
+ ML_Lex.Kind.ANTIQ_START -> LITERAL4,
+ ML_Lex.Kind.ANTIQ_STOP -> LITERAL4,
+ ML_Lex.Kind.ANTIQ_OTHER -> NULL,
+ ML_Lex.Kind.ANTIQ_STRING -> NULL,
+ ML_Lex.Kind.ANTIQ_ALT_STRING -> NULL,
+ ML_Lex.Kind.ANTIQ_CARTOUCHE -> NULL,
+ ML_Lex.Kind.ERROR -> INVALID
+ ).withDefaultValue(NULL)
+ }
+
+ def ml_token_markup(token: ML_Lex.Token): Byte =
+ if (!token.is_keyword) ml_token_style(token.kind)
+ else if (token.is_delimiter) JEditToken.OPERATOR
+ else if (ML_Lex.keywords2(token.source)) JEditToken.KEYWORD2
+ else if (ML_Lex.keywords3(token.source)) JEditToken.KEYWORD3
+ else JEditToken.KEYWORD1
+
+
+ /* markup elements */
+
+ private val indentation_elements =
+ Markup.Elements(Markup.Command_Indent.name)
+
+ private val semantic_completion_elements =
+ Markup.Elements(Markup.COMPLETION, Markup.NO_COMPLETION)
+
+ private val language_context_elements =
+ Markup.Elements(Markup.STRING, Markup.ALT_STRING, Markup.VERBATIM,
+ Markup.CARTOUCHE, Markup.COMMENT, Markup.LANGUAGE,
+ Markup.ML_STRING, Markup.ML_COMMENT)
+
+ private val language_elements = Markup.Elements(Markup.LANGUAGE)
+
+ private val citation_elements = Markup.Elements(Markup.CITATION)
+
+ private val breakpoint_elements = Markup.Elements(Markup.ML_BREAKPOINT)
+
+ private val highlight_elements =
+ Markup.Elements(Markup.EXPRESSION, Markup.CITATION, Markup.LANGUAGE, Markup.ML_TYPING,
+ Markup.TOKEN_RANGE, Markup.ENTITY, Markup.PATH, Markup.DOC, Markup.URL, Markup.SORTING,
+ Markup.TYPING, Markup.CLASS_PARAMETER, Markup.FREE, Markup.SKOLEM, Markup.BOUND,
+ Markup.VAR, Markup.TFREE, Markup.TVAR, Markup.ML_BREAKPOINT,
+ Markup.MARKDOWN_PARAGRAPH, Markup.Markdown_List.name)
+
+ private val hyperlink_elements =
+ Markup.Elements(Markup.ENTITY, Markup.PATH, Markup.DOC, Markup.POSITION,
+ Markup.CITATION, Markup.URL)
+
+ private val active_elements =
+ Markup.Elements(Markup.DIALOG, Markup.BROWSER, Markup.GRAPHVIEW,
+ Markup.SENDBACK, Markup.JEDIT_ACTION, Markup.SIMP_TRACE_PANEL)
+
+ private val tooltip_message_elements =
+ Markup.Elements(Markup.WRITELN, Markup.INFORMATION, Markup.WARNING, Markup.LEGACY, Markup.ERROR,
+ Markup.BAD)
+
+ private val gutter_elements =
+ Markup.Elements(Markup.WRITELN, Markup.INFORMATION, Markup.WARNING, Markup.LEGACY, Markup.ERROR)
+
+ private val squiggly_elements =
+ Markup.Elements(Markup.WRITELN, Markup.INFORMATION, Markup.WARNING, Markup.LEGACY, Markup.ERROR)
+
+ private val line_background_elements =
+ Markup.Elements(Markup.WRITELN_MESSAGE, Markup.STATE_MESSAGE, Markup.INFORMATION_MESSAGE,
+ Markup.TRACING_MESSAGE, Markup.WARNING_MESSAGE, Markup.LEGACY_MESSAGE,
+ Markup.ERROR_MESSAGE)
+
+ private val separator_elements =
+ Markup.Elements(Markup.SEPARATOR)
+
+ private val background_elements =
+ Protocol.proper_status_elements + Markup.WRITELN_MESSAGE +
+ Markup.STATE_MESSAGE + Markup.INFORMATION_MESSAGE +
+ Markup.TRACING_MESSAGE + Markup.WARNING_MESSAGE +
+ Markup.LEGACY_MESSAGE + Markup.ERROR_MESSAGE +
+ Markup.BAD + Markup.INTENSIFY + Markup.ENTITY +
+ Markup.Markdown_Item.name ++ active_elements
+
+ private val foreground_elements =
+ Markup.Elements(Markup.STRING, Markup.ALT_STRING, Markup.VERBATIM,
+ Markup.CARTOUCHE, Markup.ANTIQUOTED)
+
+ private val bullet_elements =
+ Markup.Elements(Markup.BULLET, Markup.ML_BREAKPOINT)
+
+ private val fold_depth_elements =
+ Markup.Elements(Markup.TEXT_FOLD, Markup.GOAL, Markup.SUBGOAL)
+}
+
+
+class JEdit_Rendering(snapshot: Document.Snapshot, options: Options)
+ extends Rendering(snapshot, options, PIDE.resources)
+{
+ /* colors */
+
+ def color_value(s: String): Color = Color_Value(options.string(s))
+
+ val outdated_color = color_value("outdated_color")
+ val unprocessed_color = color_value("unprocessed_color")
+ val unprocessed1_color = color_value("unprocessed1_color")
+ val running_color = color_value("running_color")
+ val running1_color = color_value("running1_color")
+ val bullet_color = color_value("bullet_color")
+ val tooltip_color = color_value("tooltip_color")
+ val writeln_color = color_value("writeln_color")
+ val information_color = color_value("information_color")
+ val warning_color = color_value("warning_color")
+ val legacy_color = color_value("legacy_color")
+ val error_color = color_value("error_color")
+ val writeln_message_color = color_value("writeln_message_color")
+ val information_message_color = color_value("information_message_color")
+ val tracing_message_color = color_value("tracing_message_color")
+ val warning_message_color = color_value("warning_message_color")
+ val legacy_message_color = color_value("legacy_message_color")
+ val error_message_color = color_value("error_message_color")
+ val spell_checker_color = color_value("spell_checker_color")
+ val bad_color = color_value("bad_color")
+ val intensify_color = color_value("intensify_color")
+ val entity_color = color_value("entity_color")
+ val entity_ref_color = color_value("entity_ref_color")
+ val breakpoint_disabled_color = color_value("breakpoint_disabled_color")
+ val breakpoint_enabled_color = color_value("breakpoint_enabled_color")
+ val caret_debugger_color = color_value("caret_debugger_color")
+ val quoted_color = color_value("quoted_color")
+ val antiquoted_color = color_value("antiquoted_color")
+ val antiquote_color = color_value("antiquote_color")
+ val highlight_color = color_value("highlight_color")
+ val hyperlink_color = color_value("hyperlink_color")
+ val active_color = color_value("active_color")
+ val active_hover_color = color_value("active_hover_color")
+ val active_result_color = color_value("active_result_color")
+ val keyword1_color = color_value("keyword1_color")
+ val keyword2_color = color_value("keyword2_color")
+ val keyword3_color = color_value("keyword3_color")
+ val quasi_keyword_color = color_value("quasi_keyword_color")
+ val improper_color = color_value("improper_color")
+ val operator_color = color_value("operator_color")
+ val caret_invisible_color = color_value("caret_invisible_color")
+ val completion_color = color_value("completion_color")
+ val search_color = color_value("search_color")
+
+ val tfree_color = color_value("tfree_color")
+ val tvar_color = color_value("tvar_color")
+ val free_color = color_value("free_color")
+ val skolem_color = color_value("skolem_color")
+ val bound_color = color_value("bound_color")
+ val var_color = color_value("var_color")
+ val inner_numeral_color = color_value("inner_numeral_color")
+ val inner_quoted_color = color_value("inner_quoted_color")
+ val inner_cartouche_color = color_value("inner_cartouche_color")
+ val inner_comment_color = color_value("inner_comment_color")
+ val dynamic_color = color_value("dynamic_color")
+ val class_parameter_color = color_value("class_parameter_color")
+
+ val markdown_item_color1 = color_value("markdown_item_color1")
+ val markdown_item_color2 = color_value("markdown_item_color2")
+ val markdown_item_color3 = color_value("markdown_item_color3")
+ val markdown_item_color4 = color_value("markdown_item_color4")
+
+
+ /* indentation */
+
+ def indentation(range: Text.Range): Int =
+ snapshot.select(range, JEdit_Rendering.indentation_elements, _ =>
+ {
+ case Text.Info(_, XML.Elem(Markup.Command_Indent(i), _)) => Some(i)
+ case _ => None
+ }).headOption.map(_.info).getOrElse(0)
+
+
+ /* completion */
+
+ def semantic_completion(completed_range: Option[Text.Range], range: Text.Range)
+ : Option[Text.Info[Completion.Semantic]] =
+ if (snapshot.is_outdated) None
+ else {
+ snapshot.select(range, JEdit_Rendering.semantic_completion_elements, _ =>
+ {
+ case Completion.Semantic.Info(info) =>
+ completed_range match {
+ case Some(range0) if range0.contains(info.range) && range0 != info.range => None
+ case _ => Some(info)
+ }
+ case _ => None
+ }).headOption.map(_.info)
+ }
+
+ def language_context(range: Text.Range): Option[Completion.Language_Context] =
+ snapshot.select(range, JEdit_Rendering.language_context_elements, _ =>
+ {
+ case Text.Info(_, XML.Elem(Markup.Language(language, symbols, antiquotes, delimited), _)) =>
+ if (delimited) Some(Completion.Language_Context(language, symbols, antiquotes))
+ else None
+ case Text.Info(_, elem)
+ if elem.name == Markup.ML_STRING || elem.name == Markup.ML_COMMENT =>
+ Some(Completion.Language_Context.ML_inner)
+ case Text.Info(_, _) =>
+ Some(Completion.Language_Context.inner)
+ }).headOption.map(_.info)
+
+ def language_path(range: Text.Range): Option[Text.Range] =
+ snapshot.select(range, JEdit_Rendering.language_elements, _ =>
+ {
+ case Text.Info(info_range, XML.Elem(Markup.Language(Markup.Language.PATH, _, _, _), _)) =>
+ Some(snapshot.convert(info_range))
+ case _ => None
+ }).headOption.map(_.info)
+
+ def citation(range: Text.Range): Option[Text.Info[String]] =
+ snapshot.select(range, JEdit_Rendering.citation_elements, _ =>
+ {
+ case Text.Info(info_range, XML.Elem(Markup.Citation(name), _)) =>
+ Some(Text.Info(snapshot.convert(info_range), name))
+ case _ => None
+ }).headOption.map(_.info)
+
+
+ /* spell checker */
+
+ private lazy val spell_checker_elements =
+ Markup.Elements(space_explode(',', options.string("spell_checker_elements")): _*)
+
+ def spell_checker_ranges(range: Text.Range): List[Text.Range] =
+ snapshot.select(range, spell_checker_elements, _ => _ => Some(())).map(_.range)
+
+ def spell_checker_point(range: Text.Range): Option[Text.Range] =
+ snapshot.select(range, spell_checker_elements, _ =>
+ {
+ case info => Some(snapshot.convert(info.range))
+ }).headOption.map(_.info)
+
+
+ /* breakpoints */
+
+ def breakpoint(range: Text.Range): Option[(Command, Long)] =
+ if (snapshot.is_outdated) None
+ else
+ snapshot.select(range, JEdit_Rendering.breakpoint_elements, command_states =>
+ {
+ case Text.Info(_, Protocol.ML_Breakpoint(breakpoint)) =>
+ command_states match {
+ case st :: _ => Some((st.command, breakpoint))
+ case _ => None
+ }
+ case _ => None
+ }).headOption.map(_.info)
+
+
+ /* command status overview */
+
+ def overview_color(range: Text.Range): Option[Color] =
+ {
+ if (snapshot.is_outdated) None
+ else {
+ val results =
+ snapshot.cumulate[List[Markup]](range, Nil, Protocol.liberal_status_elements, _ =>
+ {
+ case (status, Text.Info(_, elem)) => Some(elem.markup :: status)
+ }, status = true)
+ if (results.isEmpty) None
+ else {
+ val status = Protocol.Status.make(results.iterator.flatMap(_.info))
+
+ if (status.is_running) Some(running_color)
+ else if (status.is_failed) Some(error_color)
+ else if (status.is_warned) Some(warning_color)
+ else if (status.is_unprocessed) Some(unprocessed_color)
+ else None
+ }
+ }
+ }
+
+
+ /* caret focus */
+
+ def entity_ref(range: Text.Range, focus: Set[Long]): List[Text.Info[Color]] =
+ snapshot.select(range, Rendering.caret_focus_elements, _ =>
+ {
+ case Text.Info(_, XML.Elem(Markup(Markup.ENTITY, Markup.Entity.Ref(i)), _)) if focus(i) =>
+ Some(entity_ref_color)
+ case _ => None
+ })
+
+
+ /* highlighted area */
+
+ def highlight(range: Text.Range): Option[Text.Info[Color]] =
+ snapshot.select(range, JEdit_Rendering.highlight_elements, _ =>
+ {
+ case info => Some(Text.Info(snapshot.convert(info.range), highlight_color))
+ }).headOption.map(_.info)
+
+
+ /* hyperlinks */
+
+ def hyperlink(range: Text.Range): Option[Text.Info[PIDE.editor.Hyperlink]] =
+ {
+ snapshot.cumulate[Vector[Text.Info[PIDE.editor.Hyperlink]]](
+ range, Vector.empty, JEdit_Rendering.hyperlink_elements, _ =>
+ {
+ case (links, Text.Info(info_range, XML.Elem(Markup.Path(name), _))) =>
+ val file = resources.append_file(snapshot.node_name.master_dir, name)
+ val link = PIDE.editor.hyperlink_file(true, file)
+ Some(links :+ Text.Info(snapshot.convert(info_range), link))
+
+ case (links, Text.Info(info_range, XML.Elem(Markup.Doc(name), _))) =>
+ PIDE.editor.hyperlink_doc(name).map(link =>
+ (links :+ Text.Info(snapshot.convert(info_range), link)))
+
+ case (links, Text.Info(info_range, XML.Elem(Markup.Url(name), _))) =>
+ val link = PIDE.editor.hyperlink_url(name)
+ Some(links :+ Text.Info(snapshot.convert(info_range), link))
+
+ case (links, Text.Info(info_range, XML.Elem(Markup(Markup.ENTITY, props), _))) =>
+ val opt_link = PIDE.editor.hyperlink_def_position(true, snapshot, props)
+ opt_link.map(link => links :+ Text.Info(snapshot.convert(info_range), link))
+
+ case (links, Text.Info(info_range, XML.Elem(Markup(Markup.POSITION, props), _))) =>
+ val opt_link = PIDE.editor.hyperlink_position(true, snapshot, props)
+ opt_link.map(link => links :+ Text.Info(snapshot.convert(info_range), link))
+
+ case (links, Text.Info(info_range, XML.Elem(Markup.Citation(name), _))) =>
+ val opt_link =
+ Bibtex_JEdit.entries_iterator.collectFirst(
+ { case (a, buffer, offset) if a == name =>
+ PIDE.editor.hyperlink_buffer(true, buffer, offset) })
+ opt_link.map(link => links :+ Text.Info(snapshot.convert(info_range), link))
+
+ case _ => None
+ }) match { case Text.Info(_, _ :+ info) :: _ => Some(info) case _ => None }
+ }
+
+
+ /* active elements */
+
+ def active(range: Text.Range): Option[Text.Info[XML.Elem]] =
+ snapshot.select(range, JEdit_Rendering.active_elements, command_states =>
+ {
+ case Text.Info(info_range, elem) =>
+ if (elem.name == Markup.DIALOG) {
+ elem match {
+ case Protocol.Dialog(_, serial, _)
+ if !command_states.exists(st => st.results.defined(serial)) =>
+ Some(Text.Info(snapshot.convert(info_range), elem))
+ case _ => None
+ }
+ }
+ else Some(Text.Info(snapshot.convert(info_range), elem))
+ }).headOption.map(_.info)
+
+ def command_results(range: Text.Range): Command.Results =
+ Command.State.merge_results(
+ snapshot.select[List[Command.State]](range, Markup.Elements.full, command_states =>
+ { case _ => Some(command_states) }).flatMap(_.info))
+
+
+ /* tooltip messages */
+
+ def tooltip_message(range: Text.Range): Option[Text.Info[XML.Body]] =
+ {
+ val results =
+ snapshot.cumulate[List[Text.Info[Command.Results.Entry]]](
+ range, Nil, JEdit_Rendering.tooltip_message_elements, _ =>
+ {
+ case (msgs, Text.Info(info_range,
+ XML.Elem(Markup(name, props @ Markup.Serial(serial)), body)))
+ if body.nonEmpty =>
+ val entry: Command.Results.Entry =
+ serial -> XML.Elem(Markup(Markup.message(name), props), body)
+ Some(Text.Info(snapshot.convert(info_range), entry) :: msgs)
+
+ case _ => None
+ }).flatMap(_.info)
+ if (results.isEmpty) None
+ else {
+ val r = Text.Range(results.head.range.start, results.last.range.stop)
+ val msgs = Command.Results.make(results.map(_.info))
+ Some(Text.Info(r, Pretty.separate(msgs.iterator.map(_._2).toList)))
+ }
+ }
+
+
+ /* tooltips */
+
+ def tooltip_margin: Int = options.int("jedit_tooltip_margin")
+ def timing_threshold: Double = options.real("jedit_timing_threshold")
+
+ def tooltip(range: Text.Range): Option[Text.Info[XML.Body]] =
+ tooltips(range).map({ case Text.Info(r, all_tips) =>
+ Text.Info(r, Library.separate(Pretty.fbrk, all_tips)) })
+
+ lazy val tooltip_close_icon = JEdit_Lib.load_icon(options.string("tooltip_close_icon"))
+ lazy val tooltip_detach_icon = JEdit_Lib.load_icon(options.string("tooltip_detach_icon"))
+
+
+ /* gutter */
+
+ private def gutter_message_pri(msg: XML.Tree): Int =
+ if (Protocol.is_error(msg)) Rendering.error_pri
+ else if (Protocol.is_legacy(msg)) Rendering.legacy_pri
+ else if (Protocol.is_warning(msg)) Rendering.warning_pri
+ else if (Protocol.is_information(msg)) Rendering.information_pri
+ else 0
+
+ private lazy val gutter_message_content = Map(
+ Rendering.information_pri ->
+ (JEdit_Lib.load_icon(options.string("gutter_information_icon")), information_message_color),
+ Rendering.warning_pri ->
+ (JEdit_Lib.load_icon(options.string("gutter_warning_icon")), warning_message_color),
+ Rendering.legacy_pri ->
+ (JEdit_Lib.load_icon(options.string("gutter_legacy_icon")), warning_message_color),
+ Rendering.error_pri ->
+ (JEdit_Lib.load_icon(options.string("gutter_error_icon")), error_message_color))
+
+ def gutter_content(range: Text.Range): Option[(Icon, Color)] =
+ {
+ val pris =
+ snapshot.cumulate[Int](range, 0, JEdit_Rendering.gutter_elements, _ =>
+ {
+ case (pri, Text.Info(_, msg @ XML.Elem(Markup(_, Markup.Serial(serial)), _))) =>
+ Some(pri max gutter_message_pri(msg))
+ case _ => None
+ }).map(_.info)
+
+ gutter_message_content.get((0 /: pris)(_ max _))
+ }
+
+
+ /* squiggly underline */
+
+ private lazy val squiggly_colors = Map(
+ Rendering.writeln_pri -> writeln_color,
+ Rendering.information_pri -> information_color,
+ Rendering.warning_pri -> warning_color,
+ Rendering.legacy_pri -> legacy_color,
+ Rendering.error_pri -> error_color)
+
+ def squiggly_underline(range: Text.Range): List[Text.Info[Color]] =
+ {
+ val results =
+ snapshot.cumulate[Int](range, 0, JEdit_Rendering.squiggly_elements, _ =>
+ {
+ case (pri, Text.Info(_, elem)) => Some(pri max Rendering.message_pri(elem.name))
+ })
+ for {
+ Text.Info(r, pri) <- results
+ color <- squiggly_colors.get(pri)
+ } yield Text.Info(r, color)
+ }
+
+
+ /* message output */
+
+ private lazy val message_colors = Map(
+ Rendering.writeln_pri -> writeln_message_color,
+ Rendering.information_pri -> information_message_color,
+ Rendering.tracing_pri -> tracing_message_color,
+ Rendering.warning_pri -> warning_message_color,
+ Rendering.legacy_pri -> legacy_message_color,
+ Rendering.error_pri -> error_message_color)
+
+ def line_background(range: Text.Range): Option[(Color, Boolean)] =
+ {
+ val results =
+ snapshot.cumulate[Int](range, 0, JEdit_Rendering.line_background_elements, _ =>
+ {
+ case (pri, Text.Info(_, elem)) => Some(pri max Rendering.message_pri(elem.name))
+ })
+ val pri = (0 /: results) { case (p1, Text.Info(_, p2)) => p1 max p2 }
+
+ message_colors.get(pri).map(message_color =>
+ {
+ val is_separator =
+ snapshot.cumulate[Boolean](range, false, JEdit_Rendering.separator_elements, _ =>
+ {
+ case _ => Some(true)
+ }).exists(_.info)
+ (message_color, is_separator)
+ })
+ }
+
+ def output_messages(results: Command.Results): List[XML.Tree] =
+ {
+ val (states, other) =
+ results.iterator.map(_._2).filterNot(Protocol.is_result(_)).toList
+ .partition(Protocol.is_state(_))
+ states ::: other
+ }
+
+
+ /* text background */
+
+ def background(range: Text.Range, focus: Set[Long]): List[Text.Info[Color]] =
+ {
+ for {
+ Text.Info(r, result) <-
+ snapshot.cumulate[(List[Markup], Option[Color])](
+ range, (List(Markup.Empty), None), JEdit_Rendering.background_elements,
+ command_states =>
+ {
+ case (((status, color), Text.Info(_, XML.Elem(markup, _))))
+ if status.nonEmpty && Protocol.proper_status_elements(markup.name) =>
+ Some((markup :: status, color))
+ case (_, Text.Info(_, XML.Elem(Markup(Markup.BAD, _), _))) =>
+ Some((Nil, Some(bad_color)))
+ case (_, Text.Info(_, XML.Elem(Markup(Markup.INTENSIFY, _), _))) =>
+ Some((Nil, Some(intensify_color)))
+ case (_, Text.Info(_, XML.Elem(Markup(Markup.ENTITY, props), _))) =>
+ props match {
+ case Markup.Entity.Def(i) if focus(i) => Some((Nil, Some(entity_color)))
+ case Markup.Entity.Ref(i) if focus(i) => Some((Nil, Some(entity_color)))
+ case _ => None
+ }
+ case (_, Text.Info(_, XML.Elem(Markup.Markdown_Item(depth), _))) =>
+ val color =
+ depth match {
+ case 1 => markdown_item_color1
+ case 2 => markdown_item_color2
+ case 3 => markdown_item_color3
+ case _ => markdown_item_color4
+ }
+ Some((Nil, Some(color)))
+ case (acc, Text.Info(_, Protocol.Dialog(_, serial, result))) =>
+ command_states.collectFirst(
+ { case st if st.results.defined(serial) => st.results.get(serial).get }) match
+ {
+ case Some(Protocol.Dialog_Result(res)) if res == result =>
+ Some((Nil, Some(active_result_color)))
+ case _ =>
+ Some((Nil, Some(active_color)))
+ }
+ case (_, Text.Info(_, elem)) =>
+ if (JEdit_Rendering.active_elements(elem.name)) Some((Nil, Some(active_color)))
+ else None
+ })
+ color <-
+ (result match {
+ case (markups, opt_color) if markups.nonEmpty =>
+ val status = Protocol.Status.make(markups.iterator)
+ if (status.is_unprocessed) Some(unprocessed1_color)
+ else if (status.is_running) Some(running1_color)
+ else opt_color
+ case (_, opt_color) => opt_color
+ })
+ } yield Text.Info(r, color)
+ }
+
+
+ /* text foreground */
+
+ def foreground(range: Text.Range): List[Text.Info[Color]] =
+ snapshot.select(range, JEdit_Rendering.foreground_elements, _ =>
+ {
+ case Text.Info(_, elem) =>
+ if (elem.name == Markup.ANTIQUOTED) Some(antiquoted_color) else Some(quoted_color)
+ })
+
+
+ /* text color */
+
+ val foreground_color = jEdit.getColorProperty("view.fgColor")
+
+ private lazy val text_colors: Map[String, Color] = Map(
+ Markup.KEYWORD1 -> keyword1_color,
+ Markup.KEYWORD2 -> keyword2_color,
+ Markup.KEYWORD3 -> keyword3_color,
+ Markup.QUASI_KEYWORD -> quasi_keyword_color,
+ Markup.IMPROPER -> improper_color,
+ Markup.OPERATOR -> operator_color,
+ Markup.STRING -> foreground_color,
+ Markup.ALT_STRING -> foreground_color,
+ Markup.VERBATIM -> foreground_color,
+ Markup.CARTOUCHE -> foreground_color,
+ Markup.LITERAL -> keyword1_color,
+ Markup.DELIMITER -> foreground_color,
+ Markup.TFREE -> tfree_color,
+ Markup.TVAR -> tvar_color,
+ Markup.FREE -> free_color,
+ Markup.SKOLEM -> skolem_color,
+ Markup.BOUND -> bound_color,
+ Markup.VAR -> var_color,
+ Markup.INNER_STRING -> inner_quoted_color,
+ Markup.INNER_CARTOUCHE -> inner_cartouche_color,
+ Markup.INNER_COMMENT -> inner_comment_color,
+ Markup.DYNAMIC_FACT -> dynamic_color,
+ Markup.CLASS_PARAMETER -> class_parameter_color,
+ Markup.ANTIQUOTE -> antiquote_color,
+ Markup.ML_KEYWORD1 -> keyword1_color,
+ Markup.ML_KEYWORD2 -> keyword2_color,
+ Markup.ML_KEYWORD3 -> keyword3_color,
+ Markup.ML_DELIMITER -> foreground_color,
+ Markup.ML_NUMERAL -> inner_numeral_color,
+ Markup.ML_CHAR -> inner_quoted_color,
+ Markup.ML_STRING -> inner_quoted_color,
+ Markup.ML_COMMENT -> inner_comment_color,
+ Markup.SML_STRING -> inner_quoted_color,
+ Markup.SML_COMMENT -> inner_comment_color)
+
+ private lazy val text_color_elements =
+ Markup.Elements(text_colors.keySet)
+
+ def text_color(range: Text.Range, color: Color): List[Text.Info[Color]] =
+ {
+ if (color == Token_Markup.hidden_color) List(Text.Info(range, color))
+ else
+ snapshot.cumulate(range, color, text_color_elements, _ =>
+ {
+ case (_, Text.Info(_, elem)) => text_colors.get(elem.name)
+ })
+ }
+
+
+ /* virtual bullets */
+
+ def bullet(range: Text.Range): List[Text.Info[Color]] =
+ snapshot.select(range, JEdit_Rendering.bullet_elements, _ =>
+ {
+ case Text.Info(_, Protocol.ML_Breakpoint(breakpoint)) =>
+ Debugger.active_breakpoint_state(breakpoint).map(b =>
+ if (b) breakpoint_enabled_color else breakpoint_disabled_color)
+ case _ => Some(bullet_color)
+ })
+
+
+ /* text folds */
+
+ def fold_depth(range: Text.Range): List[Text.Info[Int]] =
+ snapshot.cumulate[Int](range, 0, JEdit_Rendering.fold_depth_elements, _ =>
+ {
+ case (depth, _) => Some(depth + 1)
+ })
+}
--- a/src/Tools/jEdit/src/jedit_sessions.scala Wed Jan 04 14:26:08 2017 +0100
+++ b/src/Tools/jEdit/src/jedit_sessions.scala Wed Jan 04 14:26:19 2017 +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.empty }
content.copy(known_theories =
content.known_theories.mapValues(name => name.map(File.platform_path(_))))
}
--- a/src/Tools/jEdit/src/output_dockable.scala Wed Jan 04 14:26:08 2017 +0100
+++ b/src/Tools/jEdit/src/output_dockable.scala Wed Jan 04 14:26:19 2017 +0100
@@ -66,7 +66,7 @@
val new_output =
if (!restriction.isDefined || restriction.get.contains(new_command)) {
- val rendering = Rendering(new_snapshot, PIDE.options.value)
+ val rendering = JEdit_Rendering(new_snapshot, PIDE.options.value)
rendering.output_messages(new_results)
}
else current_output
--- a/src/Tools/jEdit/src/plugin.scala Wed Jan 04 14:26:08 2017 +0100
+++ b/src/Tools/jEdit/src/plugin.scala Wed Jan 04 14:26:19 2017 +0100
@@ -158,7 +158,7 @@
}
}
- def rendering(view: View): Rendering = GUI_Thread.now
+ def rendering(view: View): JEdit_Rendering = GUI_Thread.now
{
val text_area = view.getTextArea
document_view(text_area) match {
@@ -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 ||
--- a/src/Tools/jEdit/src/pretty_text_area.scala Wed Jan 04 14:26:08 2017 +0100
+++ b/src/Tools/jEdit/src/pretty_text_area.scala Wed Jan 04 14:26:19 2017 +0100
@@ -54,10 +54,10 @@
}
private def text_rendering(base_snapshot: Document.Snapshot, base_results: Command.Results,
- formatted_body: XML.Body): (String, Rendering) =
+ formatted_body: XML.Body): (String, JEdit_Rendering) =
{
val (text, state) = document_state(base_snapshot, base_results, formatted_body)
- val rendering = Rendering(state.snapshot(), PIDE.options.value)
+ val rendering = JEdit_Rendering(state.snapshot(), PIDE.options.value)
(text, rendering)
}
}
@@ -75,7 +75,7 @@
private var current_body: XML.Body = Nil
private var current_base_snapshot = Document.Snapshot.init
private var current_base_results = Command.Results.empty
- private var current_rendering: Rendering =
+ private var current_rendering: JEdit_Rendering =
Pretty_Text_Area.text_rendering(current_base_snapshot, current_base_results, Nil)._2
private var future_refresh: Option[Future[Unit]] = None
--- a/src/Tools/jEdit/src/pretty_tooltip.scala Wed Jan 04 14:26:08 2017 +0100
+++ b/src/Tools/jEdit/src/pretty_tooltip.scala Wed Jan 04 14:26:19 2017 +0100
@@ -43,7 +43,7 @@
view: View,
parent: JComponent,
location: Point,
- rendering: Rendering,
+ rendering: JEdit_Rendering,
results: Command.Results,
info: Text.Info[XML.Body])
{
@@ -167,7 +167,7 @@
layered: JLayeredPane,
val original_parent: JComponent,
location: Point,
- rendering: Rendering,
+ rendering: JEdit_Rendering,
private val results: Command.Results,
private val info: Text.Info[XML.Body]) extends JPanel(new BorderLayout)
{
@@ -246,7 +246,7 @@
val screen = JEdit_Lib.screen_location(layered, location)
val size =
{
- val bounds = Rendering.popup_bounds
+ val bounds = JEdit_Rendering.popup_bounds
val w_max = layered.getWidth min (screen.bounds.width * bounds).toInt
val h_max = layered.getHeight min (screen.bounds.height * bounds).toInt
--- a/src/Tools/jEdit/src/rendering.scala Wed Jan 04 14:26:08 2017 +0100
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,923 +0,0 @@
-/* Title: Tools/jEdit/src/rendering.scala
- Author: Makarius
-
-Isabelle-specific implementation of quasi-abstract rendering and
-markup interpretation.
-*/
-
-package isabelle.jedit
-
-
-import isabelle._
-
-import java.awt.Color
-import javax.swing.Icon
-
-import org.gjt.sp.jedit.syntax.{Token => JEditToken}
-import org.gjt.sp.jedit.jEdit
-
-import scala.collection.immutable.SortedMap
-
-
-object Rendering
-{
- def apply(snapshot: Document.Snapshot, options: Options): Rendering =
- new Rendering(snapshot, options)
-
-
- /* message priorities */
-
- private val state_pri = 1
- private val writeln_pri = 2
- private val information_pri = 3
- private val tracing_pri = 4
- private val warning_pri = 5
- private val legacy_pri = 6
- private val error_pri = 7
-
- private val message_pri = Map(
- Markup.STATE -> state_pri,
- Markup.STATE_MESSAGE -> state_pri,
- Markup.WRITELN -> writeln_pri,
- Markup.WRITELN_MESSAGE -> writeln_pri,
- Markup.INFORMATION -> information_pri,
- Markup.INFORMATION_MESSAGE -> information_pri,
- Markup.TRACING -> tracing_pri,
- Markup.TRACING_MESSAGE -> tracing_pri,
- Markup.WARNING -> warning_pri,
- Markup.WARNING_MESSAGE -> warning_pri,
- Markup.LEGACY -> legacy_pri,
- Markup.LEGACY_MESSAGE -> legacy_pri,
- Markup.ERROR -> error_pri,
- Markup.ERROR_MESSAGE -> error_pri)
-
-
- /* popup window bounds */
-
- def popup_bounds: Double = (PIDE.options.real("jedit_popup_bounds") max 0.2) min 0.8
-
-
- /* Isabelle/Isar token markup */
-
- private val command_style: Map[String, Byte] =
- {
- import JEditToken._
- Map[String, Byte](
- Keyword.THY_END -> KEYWORD2,
- Keyword.PRF_ASM -> KEYWORD3,
- Keyword.PRF_ASM_GOAL -> KEYWORD3
- ).withDefaultValue(KEYWORD1)
- }
-
- private val token_style: Map[Token.Kind.Value, Byte] =
- {
- import JEditToken._
- Map[Token.Kind.Value, Byte](
- Token.Kind.KEYWORD -> KEYWORD2,
- Token.Kind.IDENT -> NULL,
- Token.Kind.LONG_IDENT -> NULL,
- Token.Kind.SYM_IDENT -> NULL,
- Token.Kind.VAR -> NULL,
- Token.Kind.TYPE_IDENT -> NULL,
- Token.Kind.TYPE_VAR -> NULL,
- Token.Kind.NAT -> NULL,
- Token.Kind.FLOAT -> NULL,
- Token.Kind.SPACE -> NULL,
- Token.Kind.STRING -> LITERAL1,
- Token.Kind.ALT_STRING -> LITERAL2,
- Token.Kind.VERBATIM -> COMMENT3,
- Token.Kind.CARTOUCHE -> COMMENT4,
- Token.Kind.COMMENT -> COMMENT1,
- Token.Kind.ERROR -> INVALID
- ).withDefaultValue(NULL)
- }
-
- def token_markup(syntax: Outer_Syntax, token: Token): Byte =
- if (token.is_command) command_style(syntax.keywords.kinds.getOrElse(token.content, ""))
- else if (token.is_keyword && token.source == Symbol.comment_decoded) JEditToken.NULL
- else if (token.is_delimiter) JEditToken.OPERATOR
- else token_style(token.kind)
-
-
- /* Isabelle/ML token markup */
-
- private val ml_token_style: Map[ML_Lex.Kind.Value, Byte] =
- {
- import JEditToken._
- Map[ML_Lex.Kind.Value, Byte](
- ML_Lex.Kind.KEYWORD -> NULL,
- ML_Lex.Kind.IDENT -> NULL,
- ML_Lex.Kind.LONG_IDENT -> NULL,
- ML_Lex.Kind.TYPE_VAR -> NULL,
- ML_Lex.Kind.WORD -> DIGIT,
- ML_Lex.Kind.INT -> DIGIT,
- ML_Lex.Kind.REAL -> DIGIT,
- ML_Lex.Kind.CHAR -> LITERAL2,
- ML_Lex.Kind.STRING -> LITERAL1,
- ML_Lex.Kind.SPACE -> NULL,
- ML_Lex.Kind.COMMENT -> COMMENT1,
- ML_Lex.Kind.ANTIQ -> NULL,
- ML_Lex.Kind.ANTIQ_START -> LITERAL4,
- ML_Lex.Kind.ANTIQ_STOP -> LITERAL4,
- ML_Lex.Kind.ANTIQ_OTHER -> NULL,
- ML_Lex.Kind.ANTIQ_STRING -> NULL,
- ML_Lex.Kind.ANTIQ_ALT_STRING -> NULL,
- ML_Lex.Kind.ANTIQ_CARTOUCHE -> NULL,
- ML_Lex.Kind.ERROR -> INVALID
- ).withDefaultValue(NULL)
- }
-
- def ml_token_markup(token: ML_Lex.Token): Byte =
- if (!token.is_keyword) ml_token_style(token.kind)
- else if (token.is_delimiter) JEditToken.OPERATOR
- else if (ML_Lex.keywords2(token.source)) JEditToken.KEYWORD2
- else if (ML_Lex.keywords3(token.source)) JEditToken.KEYWORD3
- else JEditToken.KEYWORD1
-
-
- /* markup elements */
-
- private val indentation_elements =
- Markup.Elements(Markup.Command_Indent.name)
-
- private val semantic_completion_elements =
- Markup.Elements(Markup.COMPLETION, Markup.NO_COMPLETION)
-
- private val language_context_elements =
- Markup.Elements(Markup.STRING, Markup.ALT_STRING, Markup.VERBATIM,
- Markup.CARTOUCHE, Markup.COMMENT, Markup.LANGUAGE,
- Markup.ML_STRING, Markup.ML_COMMENT)
-
- private val language_elements = Markup.Elements(Markup.LANGUAGE)
-
- private val citation_elements = Markup.Elements(Markup.CITATION)
-
- private val breakpoint_elements = Markup.Elements(Markup.ML_BREAKPOINT)
-
- private val caret_focus_elements = Markup.Elements(Markup.ENTITY)
-
- private val highlight_elements =
- Markup.Elements(Markup.EXPRESSION, Markup.CITATION, Markup.LANGUAGE, Markup.ML_TYPING,
- Markup.TOKEN_RANGE, Markup.ENTITY, Markup.PATH, Markup.DOC, Markup.URL, Markup.SORTING,
- Markup.TYPING, Markup.CLASS_PARAMETER, Markup.FREE, Markup.SKOLEM, Markup.BOUND,
- Markup.VAR, Markup.TFREE, Markup.TVAR, Markup.ML_BREAKPOINT,
- Markup.MARKDOWN_PARAGRAPH, Markup.Markdown_List.name)
-
- private val hyperlink_elements =
- Markup.Elements(Markup.ENTITY, Markup.PATH, Markup.DOC, Markup.POSITION,
- Markup.CITATION, Markup.URL)
-
- private val active_elements =
- Markup.Elements(Markup.DIALOG, Markup.BROWSER, Markup.GRAPHVIEW,
- Markup.SENDBACK, Markup.JEDIT_ACTION, Markup.SIMP_TRACE_PANEL)
-
- private val tooltip_message_elements =
- Markup.Elements(Markup.WRITELN, Markup.INFORMATION, Markup.WARNING, Markup.LEGACY, Markup.ERROR,
- Markup.BAD)
-
- 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 val gutter_elements =
- Markup.Elements(Markup.WRITELN, Markup.INFORMATION, Markup.WARNING, Markup.LEGACY, Markup.ERROR)
-
- private val squiggly_elements =
- Markup.Elements(Markup.WRITELN, Markup.INFORMATION, Markup.WARNING, Markup.LEGACY, Markup.ERROR)
-
- private val line_background_elements =
- Markup.Elements(Markup.WRITELN_MESSAGE, Markup.STATE_MESSAGE, Markup.INFORMATION_MESSAGE,
- Markup.TRACING_MESSAGE, Markup.WARNING_MESSAGE, Markup.LEGACY_MESSAGE,
- Markup.ERROR_MESSAGE)
-
- private val separator_elements =
- Markup.Elements(Markup.SEPARATOR)
-
- private val background_elements =
- Protocol.proper_status_elements + Markup.WRITELN_MESSAGE +
- Markup.STATE_MESSAGE + Markup.INFORMATION_MESSAGE +
- Markup.TRACING_MESSAGE + Markup.WARNING_MESSAGE +
- Markup.LEGACY_MESSAGE + Markup.ERROR_MESSAGE +
- Markup.BAD + Markup.INTENSIFY + Markup.ENTITY +
- Markup.Markdown_Item.name ++ active_elements
-
- private val foreground_elements =
- Markup.Elements(Markup.STRING, Markup.ALT_STRING, Markup.VERBATIM,
- Markup.CARTOUCHE, Markup.ANTIQUOTED)
-
- private val bullet_elements =
- Markup.Elements(Markup.BULLET, Markup.ML_BREAKPOINT)
-
- private val fold_depth_elements =
- Markup.Elements(Markup.TEXT_FOLD, Markup.GOAL, Markup.SUBGOAL)
-}
-
-
-class Rendering private(val snapshot: Document.Snapshot, val options: Options)
-{
- override def toString: String = "Rendering(" + snapshot.toString + ")"
-
-
- /* colors */
-
- def color_value(s: String): Color = Color_Value(options.string(s))
-
- val outdated_color = color_value("outdated_color")
- val unprocessed_color = color_value("unprocessed_color")
- val unprocessed1_color = color_value("unprocessed1_color")
- val running_color = color_value("running_color")
- val running1_color = color_value("running1_color")
- val bullet_color = color_value("bullet_color")
- val tooltip_color = color_value("tooltip_color")
- val writeln_color = color_value("writeln_color")
- val information_color = color_value("information_color")
- val warning_color = color_value("warning_color")
- val legacy_color = color_value("legacy_color")
- val error_color = color_value("error_color")
- val writeln_message_color = color_value("writeln_message_color")
- val information_message_color = color_value("information_message_color")
- val tracing_message_color = color_value("tracing_message_color")
- val warning_message_color = color_value("warning_message_color")
- val legacy_message_color = color_value("legacy_message_color")
- val error_message_color = color_value("error_message_color")
- val spell_checker_color = color_value("spell_checker_color")
- val bad_color = color_value("bad_color")
- val intensify_color = color_value("intensify_color")
- val entity_color = color_value("entity_color")
- val entity_ref_color = color_value("entity_ref_color")
- val breakpoint_disabled_color = color_value("breakpoint_disabled_color")
- val breakpoint_enabled_color = color_value("breakpoint_enabled_color")
- val caret_debugger_color = color_value("caret_debugger_color")
- val quoted_color = color_value("quoted_color")
- val antiquoted_color = color_value("antiquoted_color")
- val antiquote_color = color_value("antiquote_color")
- val highlight_color = color_value("highlight_color")
- val hyperlink_color = color_value("hyperlink_color")
- val active_color = color_value("active_color")
- val active_hover_color = color_value("active_hover_color")
- val active_result_color = color_value("active_result_color")
- val keyword1_color = color_value("keyword1_color")
- val keyword2_color = color_value("keyword2_color")
- val keyword3_color = color_value("keyword3_color")
- val quasi_keyword_color = color_value("quasi_keyword_color")
- val improper_color = color_value("improper_color")
- val operator_color = color_value("operator_color")
- val caret_invisible_color = color_value("caret_invisible_color")
- val completion_color = color_value("completion_color")
- val search_color = color_value("search_color")
-
- val tfree_color = color_value("tfree_color")
- val tvar_color = color_value("tvar_color")
- val free_color = color_value("free_color")
- val skolem_color = color_value("skolem_color")
- val bound_color = color_value("bound_color")
- val var_color = color_value("var_color")
- val inner_numeral_color = color_value("inner_numeral_color")
- val inner_quoted_color = color_value("inner_quoted_color")
- val inner_cartouche_color = color_value("inner_cartouche_color")
- val inner_comment_color = color_value("inner_comment_color")
- val dynamic_color = color_value("dynamic_color")
- val class_parameter_color = color_value("class_parameter_color")
-
- val markdown_item_color1 = color_value("markdown_item_color1")
- val markdown_item_color2 = color_value("markdown_item_color2")
- val markdown_item_color3 = color_value("markdown_item_color3")
- val markdown_item_color4 = color_value("markdown_item_color4")
-
-
- /* indentation */
-
- def indentation(range: Text.Range): Int =
- snapshot.select(range, Rendering.indentation_elements, _ =>
- {
- case Text.Info(_, XML.Elem(Markup.Command_Indent(i), _)) => Some(i)
- case _ => None
- }).headOption.map(_.info).getOrElse(0)
-
-
- /* completion */
-
- def semantic_completion(completed_range: Option[Text.Range], range: Text.Range)
- : Option[Text.Info[Completion.Semantic]] =
- if (snapshot.is_outdated) None
- else {
- snapshot.select(range, Rendering.semantic_completion_elements, _ =>
- {
- case Completion.Semantic.Info(info) =>
- completed_range match {
- case Some(range0) if range0.contains(info.range) && range0 != info.range => None
- case _ => Some(info)
- }
- case _ => None
- }).headOption.map(_.info)
- }
-
- def language_context(range: Text.Range): Option[Completion.Language_Context] =
- snapshot.select(range, Rendering.language_context_elements, _ =>
- {
- case Text.Info(_, XML.Elem(Markup.Language(language, symbols, antiquotes, delimited), _)) =>
- if (delimited) Some(Completion.Language_Context(language, symbols, antiquotes))
- else None
- case Text.Info(_, elem)
- if elem.name == Markup.ML_STRING || elem.name == Markup.ML_COMMENT =>
- Some(Completion.Language_Context.ML_inner)
- case Text.Info(_, _) =>
- Some(Completion.Language_Context.inner)
- }).headOption.map(_.info)
-
- def language_path(range: Text.Range): Option[Text.Range] =
- snapshot.select(range, Rendering.language_elements, _ =>
- {
- case Text.Info(info_range, XML.Elem(Markup.Language(Markup.Language.PATH, _, _, _), _)) =>
- Some(snapshot.convert(info_range))
- case _ => None
- }).headOption.map(_.info)
-
- def citation(range: Text.Range): Option[Text.Info[String]] =
- snapshot.select(range, Rendering.citation_elements, _ =>
- {
- case Text.Info(info_range, XML.Elem(Markup.Citation(name), _)) =>
- Some(Text.Info(snapshot.convert(info_range), name))
- case _ => None
- }).headOption.map(_.info)
-
-
- /* spell checker */
-
- private lazy val spell_checker_elements =
- Markup.Elements(space_explode(',', options.string("spell_checker_elements")): _*)
-
- def spell_checker_ranges(range: Text.Range): List[Text.Range] =
- snapshot.select(range, spell_checker_elements, _ => _ => Some(())).map(_.range)
-
- def spell_checker_point(range: Text.Range): Option[Text.Range] =
- snapshot.select(range, spell_checker_elements, _ =>
- {
- case info => Some(snapshot.convert(info.range))
- }).headOption.map(_.info)
-
-
- /* breakpoints */
-
- def breakpoint(range: Text.Range): Option[(Command, Long)] =
- if (snapshot.is_outdated) None
- else
- snapshot.select(range, Rendering.breakpoint_elements, command_states =>
- {
- case Text.Info(_, Protocol.ML_Breakpoint(breakpoint)) =>
- command_states match {
- case st :: _ => Some((st.command, breakpoint))
- case _ => None
- }
- case _ => None
- }).headOption.map(_.info)
-
-
- /* command status overview */
-
- def overview_color(range: Text.Range): Option[Color] =
- {
- if (snapshot.is_outdated) None
- else {
- val results =
- snapshot.cumulate[List[Markup]](range, Nil, Protocol.liberal_status_elements, _ =>
- {
- case (status, Text.Info(_, elem)) => Some(elem.markup :: status)
- }, status = true)
- if (results.isEmpty) None
- else {
- val status = Protocol.Status.make(results.iterator.flatMap(_.info))
-
- if (status.is_running) Some(running_color)
- else if (status.is_failed) Some(error_color)
- else if (status.is_warned) Some(warning_color)
- else if (status.is_unprocessed) Some(unprocessed_color)
- else None
- }
- }
- }
-
-
- /* caret focus */
-
- def entity_focus(range: Text.Range,
- check: (Boolean, Long) => Boolean = (is_def: Boolean, i: Long) => is_def): Set[Long] =
- {
- val results =
- snapshot.cumulate[Set[Long]](range, Set.empty, Rendering.caret_focus_elements, _ =>
- {
- case (serials, Text.Info(_, XML.Elem(Markup(Markup.ENTITY, props), _))) =>
- props match {
- case Markup.Entity.Def(i) if check(true, i) => Some(serials + i)
- case Markup.Entity.Ref(i) if check(false, i) => Some(serials + i)
- case _ => None
- }
- case _ => None
- })
- (Set.empty[Long] /: results){ case (s1, Text.Info(_, s2)) => s1 ++ s2 }
- }
-
- def caret_focus(caret_range: Text.Range, visible_range: Text.Range): Set[Long] =
- {
- val focus_defs = entity_focus(caret_range)
- if (focus_defs.nonEmpty) focus_defs
- else {
- val visible_defs = entity_focus(visible_range)
- entity_focus(caret_range, (is_def: Boolean, i: Long) => !is_def && visible_defs.contains(i))
- }
- }
-
- def caret_focus_ranges(caret_range: Text.Range, visible_range: Text.Range): List[Text.Range] =
- {
- val focus = caret_focus(caret_range, visible_range)
- if (focus.nonEmpty) {
- val results =
- snapshot.cumulate[Boolean](visible_range, false, Rendering.caret_focus_elements, _ =>
- {
- case (_, Text.Info(_, XML.Elem(Markup(Markup.ENTITY, props), _))) =>
- props match {
- case Markup.Entity.Def(i) if focus(i) => Some(true)
- case Markup.Entity.Ref(i) if focus(i) => Some(true)
- case _ => None
- }
- })
- for (info <- results if info.info) yield info.range
- }
- else Nil
- }
-
- def entity_ref(range: Text.Range, focus: Set[Long]): List[Text.Info[Color]] =
- snapshot.select(range, Rendering.caret_focus_elements, _ =>
- {
- case Text.Info(_, XML.Elem(Markup(Markup.ENTITY, Markup.Entity.Ref(i)), _)) if focus(i) =>
- Some(entity_ref_color)
- case _ => None
- })
-
-
- /* highlighted area */
-
- def highlight(range: Text.Range): Option[Text.Info[Color]] =
- snapshot.select(range, Rendering.highlight_elements, _ =>
- {
- case info => Some(Text.Info(snapshot.convert(info.range), highlight_color))
- }).headOption.map(_.info)
-
-
- /* hyperlinks */
-
- private def jedit_file(name: String): String =
- if (Path.is_valid(name))
- PIDE.resources.append(snapshot.node_name.master_dir, Path.explode(name))
- else name
-
- def hyperlink(range: Text.Range): Option[Text.Info[PIDE.editor.Hyperlink]] =
- {
- snapshot.cumulate[Vector[Text.Info[PIDE.editor.Hyperlink]]](
- range, Vector.empty, Rendering.hyperlink_elements, _ =>
- {
- case (links, Text.Info(info_range, XML.Elem(Markup.Path(name), _))) =>
- val link = PIDE.editor.hyperlink_file(true, jedit_file(name))
- Some(links :+ Text.Info(snapshot.convert(info_range), link))
-
- case (links, Text.Info(info_range, XML.Elem(Markup.Doc(name), _))) =>
- PIDE.editor.hyperlink_doc(name).map(link =>
- (links :+ Text.Info(snapshot.convert(info_range), link)))
-
- case (links, Text.Info(info_range, XML.Elem(Markup.Url(name), _))) =>
- val link = PIDE.editor.hyperlink_url(name)
- Some(links :+ Text.Info(snapshot.convert(info_range), link))
-
- case (links, Text.Info(info_range, XML.Elem(Markup(Markup.ENTITY, props), _)))
- if !props.exists(
- { case (Markup.KIND, Markup.ML_OPEN) => true
- case (Markup.KIND, Markup.ML_STRUCTURE) => true
- case _ => false }) =>
- val opt_link = PIDE.editor.hyperlink_def_position(true, snapshot, props)
- opt_link.map(link => links :+ Text.Info(snapshot.convert(info_range), link))
-
- case (links, Text.Info(info_range, XML.Elem(Markup(Markup.POSITION, props), _))) =>
- val opt_link = PIDE.editor.hyperlink_position(true, snapshot, props)
- opt_link.map(link => links :+ Text.Info(snapshot.convert(info_range), link))
-
- case (links, Text.Info(info_range, XML.Elem(Markup.Citation(name), _))) =>
- val opt_link =
- Bibtex_JEdit.entries_iterator.collectFirst(
- { case (a, buffer, offset) if a == name =>
- PIDE.editor.hyperlink_buffer(true, buffer, offset) })
- opt_link.map(link => links :+ Text.Info(snapshot.convert(info_range), link))
-
- case _ => None
- }) match { case Text.Info(_, _ :+ info) :: _ => Some(info) case _ => None }
- }
-
-
- /* active elements */
-
- def active(range: Text.Range): Option[Text.Info[XML.Elem]] =
- snapshot.select(range, Rendering.active_elements, command_states =>
- {
- case Text.Info(info_range, elem) =>
- if (elem.name == Markup.DIALOG) {
- elem match {
- case Protocol.Dialog(_, serial, _)
- if !command_states.exists(st => st.results.defined(serial)) =>
- Some(Text.Info(snapshot.convert(info_range), elem))
- case _ => None
- }
- }
- else Some(Text.Info(snapshot.convert(info_range), elem))
- }).headOption.map(_.info)
-
- def command_results(range: Text.Range): Command.Results =
- Command.State.merge_results(
- snapshot.select[List[Command.State]](range, Markup.Elements.full, command_states =>
- { case _ => Some(command_states) }).flatMap(_.info))
-
-
- /* tooltip messages */
-
- def tooltip_message(range: Text.Range): Option[Text.Info[XML.Body]] =
- {
- val results =
- snapshot.cumulate[List[Text.Info[Command.Results.Entry]]](
- range, Nil, Rendering.tooltip_message_elements, _ =>
- {
- case (msgs, Text.Info(info_range, msg @ XML.Elem(Markup(Markup.BAD, _), body)))
- if body.nonEmpty =>
- val entry: Command.Results.Entry = (Document_ID.make(), msg)
- Some(Text.Info(snapshot.convert(info_range), entry) :: msgs)
-
- case (msgs, Text.Info(info_range,
- XML.Elem(Markup(name, props @ Markup.Serial(serial)), body))) =>
- val entry: Command.Results.Entry =
- serial -> XML.Elem(Markup(Markup.message(name), props), body)
- Some(Text.Info(snapshot.convert(info_range), entry) :: msgs)
-
- case _ => None
- }).flatMap(_.info)
- if (results.isEmpty) None
- else {
- val r = Text.Range(results.head.range.start, results.last.range.stop)
- val msgs = Command.Results.make(results.map(_.info))
- Some(Text.Info(r, Pretty.separate(msgs.iterator.map(_._2).toList)))
- }
- }
-
-
- /* tooltips */
-
- 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(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)), Rendering.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.Path(name), _))) =>
- val file = jedit_file(name)
- val text =
- if (name == file) "file " + quote(file)
- else "path " + quote(name) + "\nfile " + quote(file)
- Some(add(prev, r, (true, XML.Text(text))))
-
- 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, Protocol.ML_Breakpoint(breakpoint))) =>
- val text =
- if (Debugger.breakpoint_state(breakpoint)) "breakpoint (enabled)"
- else "breakpoint (disabled)"
- Some(add(prev, r, (true, XML.Text(text))))
-
- 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, _), _))) =>
- Rendering.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")
-
- lazy val tooltip_close_icon = JEdit_Lib.load_icon(options.string("tooltip_close_icon"))
- lazy val tooltip_detach_icon = JEdit_Lib.load_icon(options.string("tooltip_detach_icon"))
-
-
- /* gutter */
-
- private def gutter_message_pri(msg: XML.Tree): Int =
- if (Protocol.is_error(msg)) Rendering.error_pri
- else if (Protocol.is_legacy(msg)) Rendering.legacy_pri
- else if (Protocol.is_warning(msg)) Rendering.warning_pri
- else if (Protocol.is_information(msg)) Rendering.information_pri
- else 0
-
- private lazy val gutter_message_content = Map(
- Rendering.information_pri ->
- (JEdit_Lib.load_icon(options.string("gutter_information_icon")), information_message_color),
- Rendering.warning_pri ->
- (JEdit_Lib.load_icon(options.string("gutter_warning_icon")), warning_message_color),
- Rendering.legacy_pri ->
- (JEdit_Lib.load_icon(options.string("gutter_legacy_icon")), warning_message_color),
- Rendering.error_pri ->
- (JEdit_Lib.load_icon(options.string("gutter_error_icon")), error_message_color))
-
- def gutter_content(range: Text.Range): Option[(Icon, Color)] =
- {
- val pris =
- snapshot.cumulate[Int](range, 0, Rendering.gutter_elements, _ =>
- {
- case (pri, Text.Info(_, msg @ XML.Elem(Markup(_, Markup.Serial(serial)), _))) =>
- Some(pri max gutter_message_pri(msg))
- case _ => None
- }).map(_.info)
-
- gutter_message_content.get((0 /: pris)(_ max _))
- }
-
-
- /* squiggly underline */
-
- private lazy val squiggly_colors = Map(
- Rendering.writeln_pri -> writeln_color,
- Rendering.information_pri -> information_color,
- Rendering.warning_pri -> warning_color,
- Rendering.legacy_pri -> legacy_color,
- Rendering.error_pri -> error_color)
-
- def squiggly_underline(range: Text.Range): List[Text.Info[Color]] =
- {
- val results =
- snapshot.cumulate[Int](range, 0, Rendering.squiggly_elements, _ =>
- {
- case (pri, Text.Info(_, elem)) => Some(pri max Rendering.message_pri(elem.name))
- })
- for {
- Text.Info(r, pri) <- results
- color <- squiggly_colors.get(pri)
- } yield Text.Info(r, color)
- }
-
-
- /* message output */
-
- private lazy val message_colors = Map(
- Rendering.writeln_pri -> writeln_message_color,
- Rendering.information_pri -> information_message_color,
- Rendering.tracing_pri -> tracing_message_color,
- Rendering.warning_pri -> warning_message_color,
- Rendering.legacy_pri -> legacy_message_color,
- Rendering.error_pri -> error_message_color)
-
- def line_background(range: Text.Range): Option[(Color, Boolean)] =
- {
- val results =
- snapshot.cumulate[Int](range, 0, Rendering.line_background_elements, _ =>
- {
- case (pri, Text.Info(_, elem)) => Some(pri max Rendering.message_pri(elem.name))
- })
- val pri = (0 /: results) { case (p1, Text.Info(_, p2)) => p1 max p2 }
-
- message_colors.get(pri).map(message_color =>
- {
- val is_separator =
- snapshot.cumulate[Boolean](range, false, Rendering.separator_elements, _ =>
- {
- case _ => Some(true)
- }).exists(_.info)
- (message_color, is_separator)
- })
- }
-
- def output_messages(results: Command.Results): List[XML.Tree] =
- {
- val (states, other) =
- results.iterator.map(_._2).filterNot(Protocol.is_result(_)).toList
- .partition(Protocol.is_state(_))
- states ::: other
- }
-
-
- /* text background */
-
- def background(range: Text.Range, focus: Set[Long]): List[Text.Info[Color]] =
- {
- for {
- Text.Info(r, result) <-
- snapshot.cumulate[(List[Markup], Option[Color])](
- range, (List(Markup.Empty), None), Rendering.background_elements,
- command_states =>
- {
- case (((status, color), Text.Info(_, XML.Elem(markup, _))))
- if status.nonEmpty && Protocol.proper_status_elements(markup.name) =>
- Some((markup :: status, color))
- case (_, Text.Info(_, XML.Elem(Markup(Markup.BAD, _), _))) =>
- Some((Nil, Some(bad_color)))
- case (_, Text.Info(_, XML.Elem(Markup(Markup.INTENSIFY, _), _))) =>
- Some((Nil, Some(intensify_color)))
- case (_, Text.Info(_, XML.Elem(Markup(Markup.ENTITY, props), _))) =>
- props match {
- case Markup.Entity.Def(i) if focus(i) => Some((Nil, Some(entity_color)))
- case Markup.Entity.Ref(i) if focus(i) => Some((Nil, Some(entity_color)))
- case _ => None
- }
- case (_, Text.Info(_, XML.Elem(Markup.Markdown_Item(depth), _))) =>
- val color =
- depth match {
- case 1 => markdown_item_color1
- case 2 => markdown_item_color2
- case 3 => markdown_item_color3
- case _ => markdown_item_color4
- }
- Some((Nil, Some(color)))
- case (acc, Text.Info(_, Protocol.Dialog(_, serial, result))) =>
- command_states.collectFirst(
- { case st if st.results.defined(serial) => st.results.get(serial).get }) match
- {
- case Some(Protocol.Dialog_Result(res)) if res == result =>
- Some((Nil, Some(active_result_color)))
- case _ =>
- Some((Nil, Some(active_color)))
- }
- case (_, Text.Info(_, elem)) =>
- if (Rendering.active_elements(elem.name)) Some((Nil, Some(active_color)))
- else None
- })
- color <-
- (result match {
- case (markups, opt_color) if markups.nonEmpty =>
- val status = Protocol.Status.make(markups.iterator)
- if (status.is_unprocessed) Some(unprocessed1_color)
- else if (status.is_running) Some(running1_color)
- else opt_color
- case (_, opt_color) => opt_color
- })
- } yield Text.Info(r, color)
- }
-
-
- /* text foreground */
-
- def foreground(range: Text.Range): List[Text.Info[Color]] =
- snapshot.select(range, Rendering.foreground_elements, _ =>
- {
- case Text.Info(_, elem) =>
- if (elem.name == Markup.ANTIQUOTED) Some(antiquoted_color) else Some(quoted_color)
- })
-
-
- /* text color */
-
- val foreground_color = jEdit.getColorProperty("view.fgColor")
-
- private lazy val text_colors: Map[String, Color] = Map(
- Markup.KEYWORD1 -> keyword1_color,
- Markup.KEYWORD2 -> keyword2_color,
- Markup.KEYWORD3 -> keyword3_color,
- Markup.QUASI_KEYWORD -> quasi_keyword_color,
- Markup.IMPROPER -> improper_color,
- Markup.OPERATOR -> operator_color,
- Markup.STRING -> foreground_color,
- Markup.ALT_STRING -> foreground_color,
- Markup.VERBATIM -> foreground_color,
- Markup.CARTOUCHE -> foreground_color,
- Markup.LITERAL -> keyword1_color,
- Markup.DELIMITER -> foreground_color,
- Markup.TFREE -> tfree_color,
- Markup.TVAR -> tvar_color,
- Markup.FREE -> free_color,
- Markup.SKOLEM -> skolem_color,
- Markup.BOUND -> bound_color,
- Markup.VAR -> var_color,
- Markup.INNER_STRING -> inner_quoted_color,
- Markup.INNER_CARTOUCHE -> inner_cartouche_color,
- Markup.INNER_COMMENT -> inner_comment_color,
- Markup.DYNAMIC_FACT -> dynamic_color,
- Markup.CLASS_PARAMETER -> class_parameter_color,
- Markup.ANTIQUOTE -> antiquote_color,
- Markup.ML_KEYWORD1 -> keyword1_color,
- Markup.ML_KEYWORD2 -> keyword2_color,
- Markup.ML_KEYWORD3 -> keyword3_color,
- Markup.ML_DELIMITER -> foreground_color,
- Markup.ML_NUMERAL -> inner_numeral_color,
- Markup.ML_CHAR -> inner_quoted_color,
- Markup.ML_STRING -> inner_quoted_color,
- Markup.ML_COMMENT -> inner_comment_color,
- Markup.SML_STRING -> inner_quoted_color,
- Markup.SML_COMMENT -> inner_comment_color)
-
- private lazy val text_color_elements =
- Markup.Elements(text_colors.keySet)
-
- def text_color(range: Text.Range, color: Color): List[Text.Info[Color]] =
- {
- if (color == Token_Markup.hidden_color) List(Text.Info(range, color))
- else
- snapshot.cumulate(range, color, text_color_elements, _ =>
- {
- case (_, Text.Info(_, elem)) => text_colors.get(elem.name)
- })
- }
-
-
- /* virtual bullets */
-
- def bullet(range: Text.Range): List[Text.Info[Color]] =
- snapshot.select(range, Rendering.bullet_elements, _ =>
- {
- case Text.Info(_, Protocol.ML_Breakpoint(breakpoint)) =>
- Debugger.active_breakpoint_state(breakpoint).map(b =>
- if (b) breakpoint_enabled_color else breakpoint_disabled_color)
- case _ => Some(bullet_color)
- })
-
-
- /* text folds */
-
- def fold_depth(range: Text.Range): List[Text.Info[Int]] =
- snapshot.cumulate[Int](range, 0, Rendering.fold_depth_elements, _ =>
- {
- case (depth, _) => Some(depth + 1)
- })
-}
--- a/src/Tools/jEdit/src/rich_text_area.scala Wed Jan 04 14:26:08 2017 +0100
+++ b/src/Tools/jEdit/src/rich_text_area.scala Wed Jan 04 14:26:19 2017 +0100
@@ -29,7 +29,7 @@
class Rich_Text_Area(
view: View,
text_area: TextArea,
- get_rendering: () => Rendering,
+ get_rendering: () => JEdit_Rendering,
close_action: () => Unit,
get_search_pattern: () => Option[Regex],
caret_update: () => Unit,
@@ -72,7 +72,7 @@
/* common painter state */
- @volatile private var painter_rendering: Rendering = null
+ @volatile private var painter_rendering: JEdit_Rendering = null
@volatile private var painter_clip: Shape = null
@volatile private var caret_focus: Set[Long] = Set.empty
@@ -105,7 +105,7 @@
}
}
- def robust_rendering(body: Rendering => Unit)
+ def robust_rendering(body: JEdit_Rendering => Unit)
{
robust_body(()) { body(painter_rendering) }
}
@@ -114,7 +114,7 @@
/* active areas within the text */
private class Active_Area[A](
- rendering: Rendering => Text.Range => Option[Text.Info[A]],
+ rendering: JEdit_Rendering => Text.Range => Option[Text.Info[A]],
cursor: Option[Int])
{
private var the_text_info: Option[(String, Text.Info[A])] = None
@@ -147,7 +147,7 @@
}
}
- def update_rendering(r: Rendering, range: Text.Range)
+ def update_rendering(r: JEdit_Rendering, range: Text.Range)
{ update(rendering(r)(range)) }
def reset { update(None) }
@@ -157,15 +157,15 @@
private val highlight_area =
new Active_Area[Color](
- (rendering: Rendering) => rendering.highlight _, None)
+ (rendering: JEdit_Rendering) => rendering.highlight _, None)
private val hyperlink_area =
new Active_Area[PIDE.editor.Hyperlink](
- (rendering: Rendering) => rendering.hyperlink _, Some(Cursor.HAND_CURSOR))
+ (rendering: JEdit_Rendering) => rendering.hyperlink _, Some(Cursor.HAND_CURSOR))
private val active_area =
new Active_Area[XML.Elem](
- (rendering: Rendering) => rendering.active _, Some(Cursor.DEFAULT_CURSOR))
+ (rendering: JEdit_Rendering) => rendering.active _, Some(Cursor.DEFAULT_CURSOR))
private val active_areas =
List((highlight_area, true), (hyperlink_area, true), (active_area, false))
@@ -362,7 +362,7 @@
private def caret_enabled: Boolean =
caret_visible && (!text_area.hasFocus || text_area.isCaretVisible)
- private def caret_color(rendering: Rendering, offset: Text.Offset): Color =
+ private def caret_color(rendering: JEdit_Rendering, offset: Text.Offset): Color =
{
if (text_area.isCaretVisible) text_area.getPainter.getCaretColor
else {
@@ -377,7 +377,7 @@
}
}
- private def paint_chunk_list(rendering: Rendering,
+ private def paint_chunk_list(rendering: JEdit_Rendering,
gfx: Graphics2D, line_start: Text.Offset, head: Chunk, x: Float, y: Float): Float =
{
val clip_rect = gfx.getClipBounds
--- a/src/Tools/jEdit/src/spell_checker.scala Wed Jan 04 14:26:08 2017 +0100
+++ b/src/Tools/jEdit/src/spell_checker.scala Wed Jan 04 14:26:19 2017 +0100
@@ -61,7 +61,7 @@
result.toList
}
- def current_word(text_area: TextArea, rendering: Rendering, range: Text.Range)
+ def current_word(text_area: TextArea, rendering: JEdit_Rendering, range: Text.Range)
: Option[Text.Info[String]] =
{
for {
@@ -75,7 +75,7 @@
/** completion **/
- def completion(text_area: TextArea, explicit: Boolean, rendering: Rendering)
+ def completion(text_area: TextArea, explicit: Boolean, rendering: JEdit_Rendering)
: Option[Completion.Result] =
{
for {
--- a/src/Tools/jEdit/src/text_structure.scala Wed Jan 04 14:26:08 2017 +0100
+++ b/src/Tools/jEdit/src/text_structure.scala Wed Jan 04 14:26:19 2017 +0100
@@ -86,7 +86,7 @@
val script_indent: Text.Info[Token] => Int =
{
- val opt_rendering: Option[Rendering] =
+ val opt_rendering: Option[JEdit_Rendering] =
if (PIDE.options.value.bool("jedit_indent_script"))
GUI_Thread.now {
(for {
--- a/src/Tools/jEdit/src/timing_dockable.scala Wed Jan 04 14:26:08 2017 +0100
+++ b/src/Tools/jEdit/src/timing_dockable.scala Wed Jan 04 14:26:19 2017 +0100
@@ -96,7 +96,7 @@
def print: String =
" " + Time.print_seconds(timing) + "s command " + quote(command.span.name)
def follow(snapshot: Document.Snapshot)
- { PIDE.editor.hyperlink_command(true, snapshot, command).foreach(_.follow(view)) }
+ { PIDE.editor.hyperlink_command(true, snapshot, command.id).foreach(_.follow(view)) }
}
--- a/src/Tools/jEdit/src/token_markup.scala Wed Jan 04 14:26:08 2017 +0100
+++ b/src/Tools/jEdit/src/token_markup.scala Wed Jan 04 14:26:19 2017 +0100
@@ -401,14 +401,15 @@
(line_context.context, opt_syntax) match {
case (Some(ctxt), _) if mode == "isabelle-ml" || mode == "sml" =>
val (tokens, ctxt1) = ML_Lex.tokenize_line(mode == "sml", line, ctxt)
- val styled_tokens = tokens.map(tok => (Rendering.ml_token_markup(tok), tok.source))
+ val styled_tokens =
+ tokens.map(tok => (JEdit_Rendering.ml_token_markup(tok), tok.source))
(styled_tokens, new Line_Context(line_context.mode, Some(ctxt1), structure))
case (Some(ctxt), Some(syntax)) if syntax.has_tokens =>
val (tokens, ctxt1) = Token.explode_line(syntax.keywords, line, ctxt)
val structure1 = structure.update(syntax.keywords, tokens)
val styled_tokens =
- tokens.map(tok => (Rendering.token_markup(syntax, tok), tok.source))
+ tokens.map(tok => (JEdit_Rendering.token_markup(syntax, tok), tok.source))
(styled_tokens, new Line_Context(line_context.mode, Some(ctxt1), structure1))
case _ =>
--- a/src/Tools/misc_legacy.ML Wed Jan 04 14:26:08 2017 +0100
+++ b/src/Tools/misc_legacy.ML Wed Jan 04 14:26:19 2017 +0100
@@ -201,26 +201,10 @@
(false, relift st, Thm.nprems_of st) gno state
in Seq.maps next (tacf subprems (Thm.trivial (Thm.cterm_of ctxt concl))) end;
-fun print_vars_terms ctxt n thm =
- let
- fun typed s ty = " " ^ s ^ " has type: " ^ Syntax.string_of_typ ctxt ty;
- fun find_vars (Const (c, ty)) =
- if null (Term.add_tvarsT ty []) then I
- else insert (op =) (typed c ty)
- | find_vars (Var (xi, ty)) =
- insert (op =) (typed (Term.string_of_vname xi) ty)
- | find_vars (Free _) = I
- | find_vars (Bound _) = I
- | find_vars (Abs (_, _, t)) = find_vars t
- | find_vars (t1 $ t2) = find_vars t1 #> find_vars t2;
- val prem = Logic.nth_prem (n, Thm.prop_of thm)
- val tms = find_vars prem []
- in warning (cat_lines ("Found schematic vars in assumptions:" :: tms)) end;
-
in
fun METAHYPS ctxt tacf n thm = SUBGOAL (metahyps_aux_tac ctxt tacf) n thm
- handle THM ("assume: variables", _, _) => (print_vars_terms ctxt n thm; Seq.empty)
+ handle THM ("assume: variables", _, _) => Seq.empty
end;
--- a/src/ZF/Tools/typechk.ML Wed Jan 04 14:26:08 2017 +0100
+++ b/src/ZF/Tools/typechk.ML Wed Jan 04 14:26:19 2017 +0100
@@ -120,8 +120,8 @@
"declaration of type-checking rule" #>
Method.setup @{binding typecheck}
(Method.sections
- [Args.add -- Args.colon >> K (Method.modifier TC_add @{here}),
- Args.del -- Args.colon >> K (Method.modifier TC_del @{here})]
+ [Args.add -- Args.colon >> K (Method.modifier TC_add \<^here>),
+ Args.del -- Args.colon >> K (Method.modifier TC_del \<^here>)]
>> K (fn ctxt => SIMPLE_METHOD (CHANGED (typecheck_tac ctxt))))
"ZF type-checking");