wenzelm@57491
|
1 |
Isabelle NEWS -- history of user-relevant changes
|
wenzelm@57491
|
2 |
=================================================
|
wenzelm@2553
|
3 |
|
wenzelm@62114
|
4 |
(Note: Isabelle/jEdit shows a tree-view of the NEWS file in Sidekick.)
|
wenzelm@60007
|
5 |
|
wenzelm@60331
|
6 |
|
wenzelm@62216
|
7 |
New in this Isabelle version
|
wenzelm@62216
|
8 |
----------------------------
|
wenzelm@62216
|
9 |
|
wenzelm@62440
|
10 |
*** General ***
|
wenzelm@62440
|
11 |
|
wenzelm@63121
|
12 |
* Embedded content (e.g. the inner syntax of types, terms, props) may be
|
wenzelm@63121
|
13 |
delimited uniformly via cartouches. This works better than old-fashioned
|
wenzelm@63121
|
14 |
quotes when sub-languages are nested.
|
wenzelm@63121
|
15 |
|
wenzelm@62958
|
16 |
* Type-inference improves sorts of newly introduced type variables for
|
wenzelm@62958
|
17 |
the object-logic, using its base sort (i.e. HOL.type for Isabelle/HOL).
|
wenzelm@62958
|
18 |
Thus terms like "f x" or "\<And>x. P x" without any further syntactic context
|
wenzelm@62958
|
19 |
produce x::'a::type in HOL instead of x::'a::{} in Pure. Rare
|
wenzelm@62958
|
20 |
INCOMPATIBILITY, need to provide explicit type constraints for Pure
|
wenzelm@62958
|
21 |
types where this is really intended.
|
wenzelm@62958
|
22 |
|
wenzelm@62969
|
23 |
* Simplified outer syntax: uniform category "name" includes long
|
wenzelm@62969
|
24 |
identifiers. Former "xname" / "nameref" / "name reference" has been
|
wenzelm@62969
|
25 |
discontinued.
|
wenzelm@62969
|
26 |
|
wenzelm@62807
|
27 |
* Mixfix annotations support general block properties, with syntax
|
wenzelm@62807
|
28 |
"(\<open>x=a y=b z \<dots>\<close>". Notable property names are "indent", "consistent",
|
wenzelm@62807
|
29 |
"unbreakable", "markup". The existing notation "(DIGITS" is equivalent
|
wenzelm@62807
|
30 |
to "(\<open>indent=DIGITS\<close>". The former notation "(00" for unbreakable blocks
|
wenzelm@62807
|
31 |
is superseded by "(\<open>unbreabable\<close>" --- rare INCOMPATIBILITY.
|
wenzelm@62789
|
32 |
|
wenzelm@62440
|
33 |
* New symbol \<circle>, e.g. for temporal operator.
|
wenzelm@62440
|
34 |
|
wenzelm@62453
|
35 |
* Old 'header' command is no longer supported (legacy since
|
wenzelm@62453
|
36 |
Isabelle2015).
|
wenzelm@62453
|
37 |
|
wenzelm@63273
|
38 |
* Command 'bundle' provides a local theory target to define a bundle
|
wenzelm@63273
|
39 |
from the body of specification commands (such as 'declare',
|
wenzelm@63273
|
40 |
'declaration', 'notation', 'lemmas', 'lemma'). For example:
|
wenzelm@63273
|
41 |
|
wenzelm@63273
|
42 |
bundle foo
|
wenzelm@63273
|
43 |
begin
|
wenzelm@63273
|
44 |
declare a [simp]
|
wenzelm@63273
|
45 |
declare b [intro]
|
wenzelm@63273
|
46 |
end
|
wenzelm@63272
|
47 |
|
wenzelm@63282
|
48 |
* Command 'unbundle' is like 'include', but works within a local theory
|
wenzelm@63282
|
49 |
context. Unlike "context includes ... begin", the effect of 'unbundle'
|
wenzelm@63282
|
50 |
on the target context persists, until different declarations are given.
|
wenzelm@63282
|
51 |
|
nipkow@63649
|
52 |
* Splitter in simp, auto and friends:
|
nipkow@63649
|
53 |
- The syntax "split add" has been discontinued, use plain "split".
|
nipkow@63656
|
54 |
- For situations with many conditional or case expressions,
|
nipkow@63649
|
55 |
there is an alternative splitting strategy that can be much faster.
|
nipkow@63649
|
56 |
It is selected by writing "split!" instead of "split". It applies
|
nipkow@63649
|
57 |
safe introduction and elimination rules after each split rule.
|
nipkow@63649
|
58 |
As a result the subgoal may be split into several subgoals.
|
nipkow@63649
|
59 |
|
wenzelm@63384
|
60 |
* Proof method "blast" is more robust wrt. corner cases of Pure
|
wenzelm@63384
|
61 |
statements without object-logic judgment.
|
wenzelm@63384
|
62 |
|
wenzelm@63533
|
63 |
* Pure provides basic versions of proof methods "simp" and "simp_all"
|
wenzelm@63533
|
64 |
that only know about meta-equality (==). Potential INCOMPATIBILITY in
|
wenzelm@63533
|
65 |
theory imports that merge Pure with e.g. Main of Isabelle/HOL: the order
|
wenzelm@63533
|
66 |
is relevant to avoid confusion of Pure.simp vs. HOL.simp.
|
wenzelm@63533
|
67 |
|
wenzelm@63624
|
68 |
* Commands 'prf' and 'full_prf' are somewhat more informative (again):
|
wenzelm@63821
|
69 |
proof terms are reconstructed and cleaned from administrative thm
|
wenzelm@63821
|
70 |
nodes.
|
wenzelm@63821
|
71 |
|
wenzelm@63821
|
72 |
* The command 'unfolding' and proof method "unfold" include a second
|
wenzelm@63821
|
73 |
stage where given equations are passed through the attribute "abs_def"
|
wenzelm@63821
|
74 |
before rewriting. This ensures that definitions are fully expanded,
|
wenzelm@63821
|
75 |
regardless of the actual parameters that are provided. Rare
|
wenzelm@63821
|
76 |
INCOMPATIBILITY in some corner cases: use proof method (simp only:)
|
wenzelm@63821
|
77 |
instead, or declare [[unfold_abs_def = false]] in the proof context.
|
wenzelm@63624
|
78 |
|
wenzelm@62440
|
79 |
|
wenzelm@62904
|
80 |
*** Prover IDE -- Isabelle/Scala/jEdit ***
|
wenzelm@62904
|
81 |
|
wenzelm@63135
|
82 |
* Cartouche abbreviations work both for " and ` to accomodate typical
|
wenzelm@63135
|
83 |
situations where old ASCII notation may be updated.
|
wenzelm@63135
|
84 |
|
wenzelm@63610
|
85 |
* Isabelle/ML and Standard ML files are presented in Sidekick with the
|
wenzelm@63610
|
86 |
tree structure of section headings: this special comment format is
|
wenzelm@63610
|
87 |
described in "implementation" chapter 0, e.g. (*** section ***).
|
wenzelm@63610
|
88 |
|
wenzelm@63022
|
89 |
* IDE support for the Isabelle/Pure bootstrap process, with the
|
wenzelm@63022
|
90 |
following independent stages:
|
wenzelm@63022
|
91 |
|
wenzelm@63022
|
92 |
src/Pure/ROOT0.ML
|
wenzelm@63022
|
93 |
src/Pure/ROOT.ML
|
wenzelm@63022
|
94 |
src/Pure/Pure.thy
|
wenzelm@63022
|
95 |
src/Pure/ML_Bootstrap.thy
|
wenzelm@63022
|
96 |
|
wenzelm@63022
|
97 |
The ML ROOT files act like quasi-theories in the context of theory
|
wenzelm@63022
|
98 |
ML_Bootstrap: this allows continuous checking of all loaded ML files.
|
wenzelm@63022
|
99 |
The theory files are presented with a modified header to import Pure
|
wenzelm@63022
|
100 |
from the running Isabelle instance. Results from changed versions of
|
wenzelm@63022
|
101 |
each stage are *not* propagated to the next stage, and isolated from the
|
wenzelm@63022
|
102 |
actual Isabelle/Pure that runs the IDE itself. The sequential
|
wenzelm@63308
|
103 |
dependencies of the above files are only observed for batch build.
|
wenzelm@62904
|
104 |
|
wenzelm@62987
|
105 |
* Highlighting of entity def/ref positions wrt. cursor.
|
wenzelm@62987
|
106 |
|
wenzelm@63461
|
107 |
* Refined folding mode "isabelle" based on Isar syntax: 'next' and 'qed'
|
wenzelm@63592
|
108 |
are treated as delimiters for fold structure; 'begin' and 'end'
|
wenzelm@63592
|
109 |
structure of theory specifications is treated as well.
|
wenzelm@63461
|
110 |
|
wenzelm@63608
|
111 |
* Sidekick parser "isabelle-context" shows nesting of context blocks
|
wenzelm@63608
|
112 |
according to 'begin' and 'end' structure.
|
wenzelm@63608
|
113 |
|
wenzelm@63474
|
114 |
* Syntactic indentation according to Isabelle outer syntax. Action
|
wenzelm@63474
|
115 |
"indent-lines" (shortcut C+i) indents the current line according to
|
wenzelm@63474
|
116 |
command keywords and some command substructure. Action
|
wenzelm@63455
|
117 |
"isabelle.newline" (shortcut ENTER) indents the old and the new line
|
wenzelm@63455
|
118 |
according to command keywords only; see also option
|
wenzelm@63455
|
119 |
"jedit_indent_newline".
|
wenzelm@63452
|
120 |
|
wenzelm@63474
|
121 |
* Semantic indentation for unstructured proof scripts ('apply' etc.) via
|
wenzelm@63474
|
122 |
number of subgoals. This requires information of ongoing document
|
wenzelm@63474
|
123 |
processing and may thus lag behind, when the user is editing too
|
wenzelm@63474
|
124 |
quickly; see also option "jedit_script_indent" and
|
wenzelm@63474
|
125 |
"jedit_script_indent_limit".
|
wenzelm@63474
|
126 |
|
wenzelm@63236
|
127 |
* Action "isabelle.select-entity" (shortcut CS+ENTER) selects all
|
wenzelm@63236
|
128 |
occurences of the formal entity at the caret position. This facilitates
|
wenzelm@63236
|
129 |
systematic renaming.
|
wenzelm@63236
|
130 |
|
wenzelm@63753
|
131 |
* Action "isabelle.keymap-merge" asks the user to resolve pending
|
wenzelm@63753
|
132 |
Isabelle keymap changes that are in conflict with the current jEdit
|
wenzelm@63753
|
133 |
keymap; non-conflicting changes are always applied implicitly. This
|
wenzelm@63753
|
134 |
action is automatically invoked on Isabelle/jEdit startup and thus
|
wenzelm@63753
|
135 |
increases chances that users see new keyboard shortcuts when re-using
|
wenzelm@63753
|
136 |
old keymaps.
|
wenzelm@63753
|
137 |
|
wenzelm@63032
|
138 |
* Document markup works across multiple Isar commands, e.g. the results
|
wenzelm@63032
|
139 |
established at the end of a proof are properly identified in the theorem
|
wenzelm@63032
|
140 |
statement.
|
wenzelm@63032
|
141 |
|
wenzelm@63513
|
142 |
* Command 'proof' provides information about proof outline with cases,
|
wenzelm@63513
|
143 |
e.g. for proof methods "cases", "induct", "goal_cases".
|
wenzelm@63513
|
144 |
|
wenzelm@63529
|
145 |
* Completion templates for commands involving "begin ... end" blocks,
|
wenzelm@63529
|
146 |
e.g. 'context', 'notepad'.
|
wenzelm@63529
|
147 |
|
wenzelm@63581
|
148 |
* Additional abbreviations for syntactic completion may be specified
|
wenzelm@63579
|
149 |
within the theory header as 'abbrevs', in addition to global
|
wenzelm@63581
|
150 |
$ISABELLE_HOME/etc/abbrevs and $ISABELLE_HOME_USER/etc/abbrevs as
|
wenzelm@63581
|
151 |
before. The theory syntax for 'keywords' has been simplified
|
wenzelm@63581
|
152 |
accordingly: optional abbrevs need to go into the new 'abbrevs' section.
|
wenzelm@63579
|
153 |
|
wenzelm@63675
|
154 |
* ML and document antiquotations for file-systems paths are more uniform
|
wenzelm@63675
|
155 |
and diverse:
|
wenzelm@63675
|
156 |
|
wenzelm@63675
|
157 |
@{path NAME} -- no file-system check
|
wenzelm@63675
|
158 |
@{file NAME} -- check for plain file
|
wenzelm@63675
|
159 |
@{dir NAME} -- check for directory
|
wenzelm@63675
|
160 |
|
wenzelm@63675
|
161 |
Minor INCOMPATIBILITY, former uses of @{file} and @{file_unchecked} may
|
wenzelm@63675
|
162 |
have to be changed.
|
wenzelm@63669
|
163 |
|
wenzelm@63669
|
164 |
|
wenzelm@62312
|
165 |
*** Isar ***
|
wenzelm@62312
|
166 |
|
wenzelm@63384
|
167 |
* The defining position of a literal fact \<open>prop\<close> is maintained more
|
wenzelm@63384
|
168 |
carefully, and made accessible as hyperlink in the Prover IDE.
|
wenzelm@63384
|
169 |
|
wenzelm@63384
|
170 |
* Commands 'finally' and 'ultimately' used to expose the result as
|
wenzelm@63384
|
171 |
literal fact: this accidental behaviour has been discontinued. Rare
|
wenzelm@63384
|
172 |
INCOMPATIBILITY, use more explicit means to refer to facts in Isar.
|
wenzelm@63384
|
173 |
|
wenzelm@63178
|
174 |
* Command 'axiomatization' has become more restrictive to correspond
|
wenzelm@63178
|
175 |
better to internal axioms as singleton facts with mandatory name. Minor
|
wenzelm@63178
|
176 |
INCOMPATIBILITY.
|
wenzelm@63178
|
177 |
|
wenzelm@63180
|
178 |
* Many specification elements support structured statements with 'if' /
|
wenzelm@63180
|
179 |
'for' eigen-context, e.g. 'axiomatization', 'abbreviation',
|
wenzelm@63180
|
180 |
'definition', 'inductive', 'function'.
|
wenzelm@63180
|
181 |
|
wenzelm@63094
|
182 |
* Toplevel theorem statements support eigen-context notation with 'if' /
|
wenzelm@63285
|
183 |
'for' (in postfix), which corresponds to 'assumes' / 'fixes' in the
|
wenzelm@63094
|
184 |
traditional long statement form (in prefix). Local premises are called
|
wenzelm@63094
|
185 |
"that" or "assms", respectively. Empty premises are *not* bound in the
|
wenzelm@63094
|
186 |
context: INCOMPATIBILITY.
|
wenzelm@63094
|
187 |
|
wenzelm@63043
|
188 |
* Command 'define' introduces a local (non-polymorphic) definition, with
|
wenzelm@63043
|
189 |
optional abstraction over local parameters. The syntax resembles
|
wenzelm@63047
|
190 |
'definition' and 'obtain'. It fits better into the Isar language than
|
wenzelm@63047
|
191 |
old 'def', which is now a legacy feature.
|
wenzelm@63043
|
192 |
|
wenzelm@63059
|
193 |
* Command 'obtain' supports structured statements with 'if' / 'for'
|
wenzelm@63059
|
194 |
context.
|
wenzelm@63059
|
195 |
|
wenzelm@62312
|
196 |
* Command '\<proof>' is an alias for 'sorry', with different
|
wenzelm@62312
|
197 |
typesetting. E.g. to produce proof holes in examples and documentation.
|
wenzelm@62216
|
198 |
|
wenzelm@62939
|
199 |
* The old proof method "default" has been removed (legacy since
|
wenzelm@62939
|
200 |
Isabelle2016). INCOMPATIBILITY, use "standard" instead.
|
wenzelm@62939
|
201 |
|
wenzelm@63259
|
202 |
* Proof methods may refer to the main facts via the dynamic fact
|
wenzelm@63259
|
203 |
"method_facts". This is particularly useful for Eisbach method
|
wenzelm@63259
|
204 |
definitions.
|
wenzelm@63259
|
205 |
|
wenzelm@63528
|
206 |
* Proof method "use" allows to modify the main facts of a given method
|
wenzelm@63528
|
207 |
expression, e.g.
|
wenzelm@63259
|
208 |
|
wenzelm@63259
|
209 |
(use facts in simp)
|
wenzelm@63259
|
210 |
(use facts in \<open>simp add: ...\<close>)
|
wenzelm@63259
|
211 |
|
wenzelm@62216
|
212 |
|
haftmann@63165
|
213 |
*** Pure ***
|
haftmann@63165
|
214 |
|
wenzelm@63166
|
215 |
* Code generator: config option "code_timing" triggers measurements of
|
wenzelm@63166
|
216 |
different phases of code generation. See src/HOL/ex/Code_Timing.thy for
|
wenzelm@63166
|
217 |
examples.
|
haftmann@63165
|
218 |
|
haftmann@63350
|
219 |
* Code generator: implicits in Scala (stemming from type class instances)
|
haftmann@63350
|
220 |
are generated into companion object of corresponding type class, to resolve
|
haftmann@63350
|
221 |
some situations where ambiguities may occur.
|
haftmann@63350
|
222 |
|
haftmann@63165
|
223 |
|
blanchet@62327
|
224 |
*** HOL ***
|
blanchet@62327
|
225 |
|
hoelzl@63627
|
226 |
* Renamed session HOL-Multivariate_Analysis to HOL-Analysis.
|
hoelzl@63627
|
227 |
|
hoelzl@63627
|
228 |
* Moved measure theory from HOL-Probability to HOL-Analysis. When importing
|
hoelzl@63627
|
229 |
HOL-Analysis some theorems need additional name spaces prefixes due to name
|
hoelzl@63627
|
230 |
clashes.
|
hoelzl@63627
|
231 |
INCOMPATIBILITY.
|
hoelzl@63627
|
232 |
|
eberlm@63635
|
233 |
* Number_Theory: algebraic foundation for primes: Generalisation of
|
eberlm@63635
|
234 |
predicate "prime" and introduction of predicates "prime_elem",
|
eberlm@63635
|
235 |
"irreducible", a "prime_factorization" function, and the "factorial_ring"
|
eberlm@63635
|
236 |
typeclass with instance proofs for nat, int, poly. Some theorems now have
|
eberlm@63635
|
237 |
different names, most notably "prime_def" is now "prime_nat_iff".
|
eberlm@63635
|
238 |
INCOMPATIBILITY.
|
eberlm@63552
|
239 |
|
eberlm@63552
|
240 |
* Probability: Code generation and QuickCheck for Probability Mass
|
eberlm@63552
|
241 |
Functions.
|
eberlm@63552
|
242 |
|
haftmann@63438
|
243 |
* Theory Set_Interval.thy: substantial new theorems on indexed sums
|
haftmann@63438
|
244 |
and products.
|
haftmann@63438
|
245 |
|
nipkow@63414
|
246 |
* Theory Library/LaTeXsugar.thy: New style "dummy_pats" for displaying
|
nipkow@63414
|
247 |
equations in functional programming style: variables present on the
|
nipkow@63414
|
248 |
left-hand but not on the righ-hand side are replaced by underscores.
|
nipkow@63414
|
249 |
|
haftmann@63416
|
250 |
* "surj" is a mere input abbreviation, to avoid hiding an equation in
|
haftmann@63416
|
251 |
term output. Minor INCOMPATIBILITY.
|
haftmann@63416
|
252 |
|
haftmann@63378
|
253 |
* Theory Library/Combinator_PER.thy: combinator to build partial
|
haftmann@63379
|
254 |
equivalence relations from a predicate and an equivalence relation.
|
haftmann@63378
|
255 |
|
haftmann@63376
|
256 |
* Theory Library/Perm.thy: basic facts about almost everywhere fix
|
haftmann@63376
|
257 |
bijections.
|
haftmann@63376
|
258 |
|
haftmann@63375
|
259 |
* Locale bijection establishes convenient default simp rules
|
haftmann@63375
|
260 |
like "inv f (f a) = a" for total bijections.
|
haftmann@63375
|
261 |
|
wenzelm@63343
|
262 |
* Former locale lifting_syntax is now a bundle, which is easier to
|
wenzelm@63343
|
263 |
include in a local context or theorem statement, e.g. "context includes
|
wenzelm@63343
|
264 |
lifting_syntax begin ... end". Minor INCOMPATIBILITY.
|
wenzelm@63343
|
265 |
|
haftmann@63303
|
266 |
* Code generation for scala: ambiguous implicts in class diagrams
|
haftmann@63303
|
267 |
are spelt out explicitly.
|
haftmann@63303
|
268 |
|
haftmann@63290
|
269 |
* Abstract locales semigroup, abel_semigroup, semilattice,
|
haftmann@63290
|
270 |
semilattice_neutr, ordering, ordering_top, semilattice_order,
|
haftmann@63290
|
271 |
semilattice_neutr_order, comm_monoid_set, semilattice_set,
|
haftmann@63290
|
272 |
semilattice_neutr_set, semilattice_order_set, semilattice_order_neutr_set
|
haftmann@63290
|
273 |
monoid_list, comm_monoid_list, comm_monoid_list_set, comm_monoid_mset,
|
haftmann@63290
|
274 |
comm_monoid_fun use boldified syntax uniformly that does not clash
|
haftmann@63290
|
275 |
with corresponding global syntax. INCOMPATIBILITY.
|
haftmann@63290
|
276 |
|
haftmann@63237
|
277 |
* Conventional syntax "%(). t" for unit abstractions. Slight syntactic
|
haftmann@63237
|
278 |
INCOMPATIBILITY.
|
haftmann@63237
|
279 |
|
haftmann@63174
|
280 |
* Command 'code_reflect' accepts empty constructor lists for datatypes,
|
haftmann@63174
|
281 |
which renders those abstract effectively.
|
haftmann@63174
|
282 |
|
haftmann@63175
|
283 |
* Command 'export_code' checks given constants for abstraction violations:
|
haftmann@63175
|
284 |
a small guarantee that given constants specify a safe interface for the
|
haftmann@63175
|
285 |
generated code.
|
haftmann@63175
|
286 |
|
eberlm@63144
|
287 |
* Probability/Random_Permutations.thy contains some theory about
|
eberlm@63144
|
288 |
choosing a permutation of a set uniformly at random and folding over a
|
eberlm@63144
|
289 |
list in random order.
|
eberlm@63144
|
290 |
|
Andreas@63246
|
291 |
* Probability/SPMF formalises discrete subprobability distributions.
|
Andreas@63246
|
292 |
|
wenzelm@63283
|
293 |
* Library/FinFun.thy: bundles "finfun_syntax" and "no_finfun_syntax"
|
wenzelm@63283
|
294 |
allow to control optional syntax in local contexts; this supersedes
|
wenzelm@63283
|
295 |
former Library/FinFun_Syntax.thy. INCOMPATIBILITY, e.g. use "unbundle
|
wenzelm@63283
|
296 |
finfun_syntax" to imitate import of "~~/src/HOL/Library/FinFun_Syntax".
|
wenzelm@63283
|
297 |
|
eberlm@63144
|
298 |
* Library/Set_Permutations.thy (executably) defines the set of
|
eberlm@63144
|
299 |
permutations of a set, i.e. the set of all lists that contain every
|
eberlm@63144
|
300 |
element of the carrier set exactly once.
|
eberlm@63144
|
301 |
|
haftmann@63161
|
302 |
* Static evaluators (Code_Evaluation.static_* in Isabelle/ML) rely on
|
haftmann@63161
|
303 |
explicitly provided auxiliary definitions for required type class
|
haftmann@63161
|
304 |
dictionaries rather than half-working magic. INCOMPATIBILITY, see
|
haftmann@63161
|
305 |
the tutorial on code generation for details.
|
haftmann@63161
|
306 |
|
wenzelm@62522
|
307 |
* New abbreviations for negated existence (but not bounded existence):
|
wenzelm@62522
|
308 |
|
wenzelm@62522
|
309 |
\<nexists>x. P x \<equiv> \<not> (\<exists>x. P x)
|
wenzelm@62522
|
310 |
\<nexists>!x. P x \<equiv> \<not> (\<exists>!x. P x)
|
wenzelm@62522
|
311 |
|
wenzelm@62521
|
312 |
* The print mode "HOL" for ASCII syntax of binders "!", "?", "?!", "@"
|
wenzelm@62521
|
313 |
has been removed for output. It is retained for input only, until it is
|
wenzelm@62521
|
314 |
eliminated altogether.
|
wenzelm@62521
|
315 |
|
blanchet@63785
|
316 |
* metis: The problem encoding has changed very slightly. This might
|
blanchet@63785
|
317 |
break existing proofs. INCOMPATIBILITY.
|
blanchet@63785
|
318 |
|
blanchet@63116
|
319 |
* Sledgehammer:
|
blanchet@63699
|
320 |
- The MaSh relevance filter has been sped up.
|
blanchet@63116
|
321 |
- Produce syntactically correct Vampire 4.0 problem files.
|
blanchet@63116
|
322 |
|
blanchet@62327
|
323 |
* (Co)datatype package:
|
blanchet@62693
|
324 |
- New commands for defining corecursive functions and reasoning about
|
blanchet@62693
|
325 |
them in "~~/src/HOL/Library/BNF_Corec.thy": 'corec', 'corecursive',
|
blanchet@62693
|
326 |
'friend_of_corec', and 'corecursion_upto'; and 'corec_unique' proof
|
blanchet@62842
|
327 |
method. See 'isabelle doc corec'.
|
blanchet@62693
|
328 |
- The predicator :: ('a => bool) => 'a F => bool is now a first-class
|
blanchet@63855
|
329 |
citizen in bounded natural functors.
|
blanchet@62693
|
330 |
- 'primrec' now allows nested calls through the predicator in addition
|
blanchet@62327
|
331 |
to the map function.
|
blanchet@63855
|
332 |
- 'bnf' automatically discharges reflexive proof obligations.
|
blanchet@62693
|
333 |
- 'bnf' outputs a slightly modified proof obligation expressing rel in
|
traytel@62332
|
334 |
terms of map and set
|
blanchet@63855
|
335 |
(not giving a specification for rel makes this one reflexive).
|
blanchet@62693
|
336 |
- 'bnf' outputs a new proof obligation expressing pred in terms of set
|
blanchet@63855
|
337 |
(not giving a specification for pred makes this one reflexive).
|
blanchet@63855
|
338 |
INCOMPATIBILITY: manual 'bnf' declarations may need adjustment.
|
blanchet@62335
|
339 |
- Renamed lemmas:
|
blanchet@62335
|
340 |
rel_prod_apply ~> rel_prod_inject
|
blanchet@62335
|
341 |
pred_prod_apply ~> pred_prod_inject
|
blanchet@62335
|
342 |
INCOMPATIBILITY.
|
blanchet@62536
|
343 |
- The "size" plugin has been made compatible again with locales.
|
blanchet@63855
|
344 |
- The theorems about "rel" and "set" may have a slightly different (but
|
blanchet@63855
|
345 |
equivalent) form.
|
blanchet@63855
|
346 |
INCOMPATIBILITY.
|
blanchet@62327
|
347 |
|
wenzelm@63807
|
348 |
* Some old / obsolete theorems have been renamed / removed, potential
|
wenzelm@63807
|
349 |
INCOMPATIBILITY.
|
wenzelm@63807
|
350 |
|
wenzelm@63807
|
351 |
nat_less_cases -- removed, use linorder_cases instead
|
wenzelm@63807
|
352 |
inv_image_comp -- removed, use image_inv_f_f instead
|
wenzelm@63807
|
353 |
image_surj_f_inv_f ~> image_f_inv_f
|
wenzelm@63113
|
354 |
|
Mathias@63456
|
355 |
* Some theorems about groups and orders have been generalised from
|
Mathias@63456
|
356 |
groups to semi-groups that are also monoids:
|
Mathias@63456
|
357 |
le_add_same_cancel1
|
Mathias@63456
|
358 |
le_add_same_cancel2
|
Mathias@63456
|
359 |
less_add_same_cancel1
|
Mathias@63456
|
360 |
less_add_same_cancel2
|
Mathias@63456
|
361 |
add_le_same_cancel1
|
Mathias@63456
|
362 |
add_le_same_cancel2
|
Mathias@63456
|
363 |
add_less_same_cancel1
|
Mathias@63456
|
364 |
add_less_same_cancel2
|
Mathias@63456
|
365 |
|
Mathias@63456
|
366 |
* Some simplifications theorems about rings have been removed, since
|
Mathias@63456
|
367 |
superseeded by a more general version:
|
Mathias@63456
|
368 |
less_add_cancel_left_greater_zero ~> less_add_same_cancel1
|
Mathias@63456
|
369 |
less_add_cancel_right_greater_zero ~> less_add_same_cancel2
|
Mathias@63456
|
370 |
less_eq_add_cancel_left_greater_eq_zero ~> le_add_same_cancel1
|
Mathias@63456
|
371 |
less_eq_add_cancel_right_greater_eq_zero ~> le_add_same_cancel2
|
Mathias@63456
|
372 |
less_eq_add_cancel_left_less_eq_zero ~> add_le_same_cancel1
|
Mathias@63456
|
373 |
less_eq_add_cancel_right_less_eq_zero ~> add_le_same_cancel2
|
Mathias@63456
|
374 |
less_add_cancel_left_less_zero ~> add_less_same_cancel1
|
Mathias@63456
|
375 |
less_add_cancel_right_less_zero ~> add_less_same_cancel2
|
Mathias@63456
|
376 |
INCOMPATIBILITY.
|
Mathias@63456
|
377 |
|
wenzelm@62407
|
378 |
* Renamed split_if -> if_split and split_if_asm -> if_split_asm to
|
wenzelm@62407
|
379 |
resemble the f.split naming convention, INCOMPATIBILITY.
|
nipkow@62396
|
380 |
|
haftmann@62597
|
381 |
* Characters (type char) are modelled as finite algebraic type
|
haftmann@62597
|
382 |
corresponding to {0..255}.
|
haftmann@62597
|
383 |
|
haftmann@62597
|
384 |
- Logical representation:
|
haftmann@62597
|
385 |
* 0 is instantiated to the ASCII zero character.
|
wenzelm@62645
|
386 |
* All other characters are represented as "Char n"
|
haftmann@62597
|
387 |
with n being a raw numeral expression less than 256.
|
wenzelm@62645
|
388 |
* Expressions of the form "Char n" with n greater than 255
|
haftmann@62597
|
389 |
are non-canonical.
|
haftmann@62597
|
390 |
- Printing and parsing:
|
wenzelm@62645
|
391 |
* Printable characters are printed and parsed as "CHR ''\<dots>''"
|
haftmann@62597
|
392 |
(as before).
|
wenzelm@62645
|
393 |
* The ASCII zero character is printed and parsed as "0".
|
haftmann@62678
|
394 |
* All other canonical characters are printed as "CHR 0xXX"
|
haftmann@62678
|
395 |
with XX being the hexadecimal character code. "CHR n"
|
haftmann@62597
|
396 |
is parsable for every numeral expression n.
|
haftmann@62598
|
397 |
* Non-canonical characters have no special syntax and are
|
haftmann@62597
|
398 |
printed as their logical representation.
|
haftmann@62597
|
399 |
- Explicit conversions from and to the natural numbers are
|
haftmann@62597
|
400 |
provided as char_of_nat, nat_of_char (as before).
|
haftmann@62597
|
401 |
- The auxiliary nibble type has been discontinued.
|
haftmann@62597
|
402 |
|
haftmann@62597
|
403 |
INCOMPATIBILITY.
|
haftmann@62597
|
404 |
|
haftmann@62430
|
405 |
* Multiset membership is now expressed using set_mset rather than count.
|
haftmann@62430
|
406 |
|
haftmann@62430
|
407 |
- Expressions "count M a > 0" and similar simplify to membership
|
haftmann@62430
|
408 |
by default.
|
haftmann@62430
|
409 |
|
haftmann@62430
|
410 |
- Converting between "count M a = 0" and non-membership happens using
|
haftmann@62430
|
411 |
equations count_eq_zero_iff and not_in_iff.
|
haftmann@62430
|
412 |
|
haftmann@62430
|
413 |
- Rules count_inI and in_countE obtain facts of the form
|
haftmann@62430
|
414 |
"count M a = n" from membership.
|
haftmann@62430
|
415 |
|
haftmann@62430
|
416 |
- Rules count_in_diffI and in_diff_countE obtain facts of the form
|
haftmann@62430
|
417 |
"count M a = n + count N a" from membership on difference sets.
|
haftmann@62430
|
418 |
|
haftmann@62430
|
419 |
INCOMPATIBILITY.
|
haftmann@62430
|
420 |
|
Mathias@63312
|
421 |
* The names of multiset theorems have been normalised to distinguish which
|
Mathias@63312
|
422 |
ordering the theorems are about
|
Mathias@63312
|
423 |
mset_less_eqI ~> mset_subset_eqI
|
Mathias@63312
|
424 |
mset_less_insertD ~> mset_subset_insertD
|
Mathias@63312
|
425 |
mset_less_eq_count ~> mset_subset_eq_count
|
Mathias@63312
|
426 |
mset_less_diff_self ~> mset_subset_diff_self
|
Mathias@63312
|
427 |
mset_le_exists_conv ~> mset_subset_eq_exists_conv
|
Mathias@63312
|
428 |
mset_le_mono_add_right_cancel ~> mset_subset_eq_mono_add_right_cancel
|
Mathias@63312
|
429 |
mset_le_mono_add_left_cancel ~> mset_subset_eq_mono_add_left_cancel
|
Mathias@63312
|
430 |
mset_le_mono_add ~> mset_subset_eq_mono_add
|
Mathias@63312
|
431 |
mset_le_add_left ~> mset_subset_eq_add_left
|
Mathias@63312
|
432 |
mset_le_add_right ~> mset_subset_eq_add_right
|
Mathias@63312
|
433 |
mset_le_single ~> mset_subset_eq_single
|
Mathias@63312
|
434 |
mset_le_multiset_union_diff_commute ~> mset_subset_eq_multiset_union_diff_commute
|
Mathias@63312
|
435 |
diff_le_self ~> diff_subset_eq_self
|
Mathias@63312
|
436 |
mset_leD ~> mset_subset_eqD
|
Mathias@63312
|
437 |
mset_lessD ~> mset_subsetD
|
Mathias@63312
|
438 |
mset_le_insertD ~> mset_subset_eq_insertD
|
Mathias@63312
|
439 |
mset_less_of_empty ~> mset_subset_of_empty
|
Mathias@63312
|
440 |
le_empty ~> subset_eq_empty
|
Mathias@63312
|
441 |
mset_less_add_bothsides ~> mset_subset_add_bothsides
|
Mathias@63312
|
442 |
mset_less_empty_nonempty ~> mset_subset_empty_nonempty
|
Mathias@63312
|
443 |
mset_less_size ~> mset_subset_size
|
Mathias@63312
|
444 |
wf_less_mset_rel ~> wf_subset_mset_rel
|
Mathias@63312
|
445 |
count_le_replicate_mset_le ~> count_le_replicate_mset_subset_eq
|
Mathias@63312
|
446 |
mset_remdups_le ~> mset_remdups_subset_eq
|
Mathias@63312
|
447 |
ms_lesseq_impl ~> subset_eq_mset_impl
|
Mathias@63312
|
448 |
|
Mathias@63312
|
449 |
Some functions have been renamed:
|
Mathias@63312
|
450 |
ms_lesseq_impl -> subset_eq_mset_impl
|
Mathias@63312
|
451 |
|
Mathias@63388
|
452 |
* Multisets are now ordered with the multiset ordering
|
Mathias@63388
|
453 |
#\<subseteq># ~> \<le>
|
Mathias@63388
|
454 |
#\<subset># ~> <
|
Mathias@63388
|
455 |
le_multiset ~> less_eq_multiset
|
Mathias@63388
|
456 |
less_multiset ~> le_multiset
|
blanchet@63407
|
457 |
INCOMPATIBILITY.
|
Mathias@63388
|
458 |
|
Mathias@63388
|
459 |
* The prefix multiset_order has been discontinued: the theorems can be directly
|
blanchet@63407
|
460 |
accessed. As a consequence, the lemmas "order_multiset" and "linorder_multiset"
|
blanchet@63407
|
461 |
have been discontinued, and the interpretations "multiset_linorder" and
|
blanchet@63407
|
462 |
"multiset_wellorder" have been replaced by instantiations.
|
blanchet@63407
|
463 |
INCOMPATIBILITY.
|
Mathias@63388
|
464 |
|
Mathias@63388
|
465 |
* Some theorems about the multiset ordering have been renamed:
|
Mathias@63388
|
466 |
le_multiset_def ~> less_eq_multiset_def
|
Mathias@63388
|
467 |
less_multiset_def ~> le_multiset_def
|
Mathias@63388
|
468 |
less_eq_imp_le_multiset ~> subset_eq_imp_le_multiset
|
Mathias@63388
|
469 |
mult_less_not_refl ~> mset_le_not_refl
|
Mathias@63388
|
470 |
mult_less_trans ~> mset_le_trans
|
Mathias@63388
|
471 |
mult_less_not_sym ~> mset_le_not_sym
|
Mathias@63388
|
472 |
mult_less_asym ~> mset_le_asym
|
Mathias@63388
|
473 |
mult_less_irrefl ~> mset_le_irrefl
|
Mathias@63388
|
474 |
union_less_mono2{,1,2} ~> union_le_mono2{,1,2}
|
Mathias@63388
|
475 |
|
Mathias@63388
|
476 |
le_multiset\<^sub>H\<^sub>O ~> less_eq_multiset\<^sub>H\<^sub>O
|
Mathias@63388
|
477 |
le_multiset_total ~> less_eq_multiset_total
|
Mathias@63388
|
478 |
less_multiset_right_total ~> subset_eq_imp_le_multiset
|
Mathias@63388
|
479 |
le_multiset_empty_left ~> less_eq_multiset_empty_left
|
Mathias@63388
|
480 |
le_multiset_empty_right ~> less_eq_multiset_empty_right
|
Mathias@63388
|
481 |
less_multiset_empty_right ~> le_multiset_empty_left
|
Mathias@63388
|
482 |
less_multiset_empty_left ~> le_multiset_empty_right
|
Mathias@63388
|
483 |
union_less_diff_plus ~> union_le_diff_plus
|
Mathias@63388
|
484 |
ex_gt_count_imp_less_multiset ~> ex_gt_count_imp_le_multiset
|
Mathias@63388
|
485 |
less_multiset_plus_left_nonempty ~> le_multiset_plus_left_nonempty
|
Mathias@63388
|
486 |
le_multiset_plus_right_nonempty ~> le_multiset_plus_right_nonempty
|
Mathias@63388
|
487 |
less_multiset_plus_plus_left_iff ~> le_multiset_plus_plus_left_iff
|
Mathias@63388
|
488 |
less_multiset_plus_plus_right_iff ~> le_multiset_plus_plus_right_iff
|
blanchet@63407
|
489 |
INCOMPATIBILITY.
|
Mathias@63388
|
490 |
|
Mathias@63524
|
491 |
* The lemma mset_map has now the attribute [simp].
|
Mathias@63524
|
492 |
INCOMPATIBILITY.
|
Mathias@63524
|
493 |
|
Mathias@63525
|
494 |
* Some theorems about multisets have been removed:
|
Mathias@63525
|
495 |
le_multiset_plus_plus_left_iff ~> add_less_cancel_right
|
Mathias@63525
|
496 |
le_multiset_plus_plus_right_iff ~> add_less_cancel_left
|
Mathias@63525
|
497 |
add_eq_self_empty_iff ~> add_cancel_left_right
|
Mathias@63793
|
498 |
mset_subset_add_bothsides ~> subset_mset.add_less_cancel_right
|
Mathias@63525
|
499 |
INCOMPATIBILITY.
|
Mathias@63525
|
500 |
|
Mathias@63410
|
501 |
* Some typeclass constraints about multisets have been reduced from ordered or
|
Mathias@63410
|
502 |
linordered to preorder. Multisets have the additional typeclasses order_bot,
|
Mathias@63410
|
503 |
no_top, ordered_ab_semigroup_add_imp_le, ordered_cancel_comm_monoid_add,
|
Mathias@63525
|
504 |
linordered_cancel_ab_semigroup_add, and ordered_ab_semigroup_monoid_add_imp_le.
|
Mathias@63410
|
505 |
INCOMPATIBILITY.
|
Mathias@63410
|
506 |
|
Mathias@63561
|
507 |
* There are some new simplification rules about multisets, the multiset
|
Mathias@63561
|
508 |
ordering, and the subset ordering on multisets.
|
Mathias@63561
|
509 |
INCOMPATIBILITY.
|
Mathias@63561
|
510 |
|
Mathias@63795
|
511 |
* The subset ordering on multisets has now the interpretations
|
Mathias@63795
|
512 |
ordered_ab_semigroup_monoid_add_imp_le and bounded_lattice_bot.
|
Mathias@63410
|
513 |
INCOMPATIBILITY.
|
Mathias@63410
|
514 |
|
Mathias@63793
|
515 |
* Multiset: single has been removed in favor of add_mset that roughly
|
Mathias@63793
|
516 |
corresponds to Set.insert. Some theorems have removed or changed:
|
Mathias@63793
|
517 |
single_not_empty ~> add_mset_not_empty or empty_not_add_mset
|
Mathias@63793
|
518 |
fold_mset_insert ~> fold_mset_add_mset
|
Mathias@63793
|
519 |
image_mset_insert ~> image_mset_add_mset
|
Mathias@63793
|
520 |
union_single_eq_diff
|
Mathias@63793
|
521 |
multi_self_add_other_not_self
|
Mathias@63793
|
522 |
diff_single_eq_union
|
Mathias@63793
|
523 |
INCOMPATIBILITY.
|
Mathias@63793
|
524 |
|
Mathias@63793
|
525 |
* Multiset: some theorems have been changed to use add_mset instead of single:
|
Mathias@63793
|
526 |
mset_add
|
Mathias@63793
|
527 |
multi_self_add_other_not_self
|
Mathias@63793
|
528 |
diff_single_eq_union
|
Mathias@63793
|
529 |
union_single_eq_diff
|
Mathias@63793
|
530 |
union_single_eq_member
|
Mathias@63793
|
531 |
add_eq_conv_diff
|
Mathias@63793
|
532 |
insert_noteq_member
|
Mathias@63793
|
533 |
add_eq_conv_ex
|
Mathias@63793
|
534 |
multi_member_split
|
Mathias@63793
|
535 |
multiset_add_sub_el_shuffle
|
Mathias@63793
|
536 |
mset_subset_eq_insertD
|
Mathias@63793
|
537 |
mset_subset_insertD
|
Mathias@63793
|
538 |
insert_subset_eq_iff
|
Mathias@63793
|
539 |
insert_union_subset_iff
|
Mathias@63793
|
540 |
multi_psub_of_add_self
|
Mathias@63793
|
541 |
inter_add_left1
|
Mathias@63793
|
542 |
inter_add_left2
|
Mathias@63793
|
543 |
inter_add_right1
|
Mathias@63793
|
544 |
inter_add_right2
|
Mathias@63793
|
545 |
sup_union_left1
|
Mathias@63793
|
546 |
sup_union_left2
|
Mathias@63793
|
547 |
sup_union_right1
|
Mathias@63793
|
548 |
sup_union_right2
|
Mathias@63793
|
549 |
size_eq_Suc_imp_eq_union
|
Mathias@63793
|
550 |
multi_nonempty_split
|
Mathias@63793
|
551 |
mset_insort
|
Mathias@63793
|
552 |
mset_update
|
Mathias@63793
|
553 |
mult1I
|
Mathias@63793
|
554 |
less_add
|
Mathias@63793
|
555 |
mset_zip_take_Cons_drop_twice
|
Mathias@63793
|
556 |
rel_mset_Zero
|
Mathias@63793
|
557 |
msed_map_invL
|
Mathias@63793
|
558 |
msed_map_invR
|
Mathias@63793
|
559 |
msed_rel_invL
|
Mathias@63793
|
560 |
msed_rel_invR
|
Mathias@63793
|
561 |
le_multiset_right_total
|
Mathias@63793
|
562 |
multiset_induct
|
Mathias@63793
|
563 |
multiset_induct2_size
|
Mathias@63793
|
564 |
multiset_induct2
|
Mathias@63793
|
565 |
INCOMPATIBILITY.
|
Mathias@63793
|
566 |
|
Mathias@63793
|
567 |
* Multiset: the definitions of some constants have changed to use add_mset instead
|
Mathias@63793
|
568 |
of adding a single element:
|
Mathias@63793
|
569 |
image_mset
|
Mathias@63793
|
570 |
mset
|
Mathias@63793
|
571 |
replicate_mset
|
Mathias@63793
|
572 |
mult1
|
Mathias@63793
|
573 |
pred_mset
|
Mathias@63793
|
574 |
rel_mset'
|
Mathias@63793
|
575 |
mset_insort
|
Mathias@63793
|
576 |
INCOMPATIBILITY.
|
Mathias@63793
|
577 |
|
Mathias@63793
|
578 |
* Due to the above changes, the attributes of some multiset theorems have
|
Mathias@63793
|
579 |
been changed:
|
Mathias@63793
|
580 |
insert_DiffM [] ~> [simp]
|
Mathias@63793
|
581 |
insert_DiffM2 [simp] ~> []
|
Mathias@63793
|
582 |
diff_add_mset_swap [simp]
|
Mathias@63793
|
583 |
fold_mset_add_mset [simp]
|
Mathias@63793
|
584 |
diff_diff_add [simp] (for multisets only)
|
Mathias@63793
|
585 |
diff_cancel [simp] ~> []
|
Mathias@63793
|
586 |
count_single [simp] ~> []
|
Mathias@63793
|
587 |
set_mset_single [simp] ~> []
|
Mathias@63793
|
588 |
size_multiset_single [simp] ~> []
|
Mathias@63793
|
589 |
size_single [simp] ~> []
|
Mathias@63793
|
590 |
image_mset_single [simp] ~> []
|
Mathias@63793
|
591 |
mset_subset_eq_mono_add_right_cancel [simp] ~> []
|
Mathias@63793
|
592 |
mset_subset_eq_mono_add_left_cancel [simp] ~> []
|
Mathias@63793
|
593 |
fold_mset_single [simp] ~> []
|
Mathias@63793
|
594 |
subset_eq_empty [simp] ~> []
|
Mathias@63795
|
595 |
empty_sup [simp] ~> []
|
Mathias@63795
|
596 |
sup_empty [simp] ~> []
|
Mathias@63795
|
597 |
inter_empty [simp] ~> []
|
Mathias@63795
|
598 |
empty_inter [simp] ~> []
|
Mathias@63793
|
599 |
INCOMPATIBILITY.
|
Mathias@63793
|
600 |
|
Mathias@63793
|
601 |
* The order of the variables in the second cases of multiset_induct,
|
Mathias@63793
|
602 |
multiset_induct2_size, multiset_induct2 has been changed (e.g. Add A a ~> Add a A).
|
Mathias@63793
|
603 |
INCOMPATIBILITY.
|
Mathias@63793
|
604 |
|
Mathias@63793
|
605 |
* There is now a simplification procedure on multisets. It mimics the behavior
|
Mathias@63793
|
606 |
of the procedure on natural numbers.
|
Mathias@63793
|
607 |
INCOMPATIBILITY.
|
Mathias@63793
|
608 |
|
nipkow@63830
|
609 |
* Renamed sums and products of multisets:
|
nipkow@63830
|
610 |
msetsum ~> sum_mset
|
nipkow@63830
|
611 |
msetprod ~> prod_mset
|
nipkow@63830
|
612 |
|
Mathias@63795
|
613 |
* The lemma one_step_implies_mult_aux on multisets has been removed, use
|
Mathias@63795
|
614 |
one_step_implies_mult instead.
|
Mathias@63795
|
615 |
INCOMPATIBILITY.
|
Mathias@63795
|
616 |
|
haftmann@62343
|
617 |
* Compound constants INFIMUM and SUPREMUM are mere abbreviations now.
|
haftmann@62343
|
618 |
INCOMPATIBILITY.
|
haftmann@62343
|
619 |
|
lp15@62408
|
620 |
* More complex analysis including Cauchy's inequality, Liouville theorem,
|
lp15@63078
|
621 |
open mapping theorem, maximum modulus principle, Residue theorem, Schwarz Lemma.
|
lp15@63078
|
622 |
|
lp15@63078
|
623 |
* Theory of polyhedra: faces, extreme points, polytopes, and the Krein–Milman
|
lp15@63078
|
624 |
Minkowski theorem.
|
lp15@62408
|
625 |
|
haftmann@62358
|
626 |
* "Gcd (f ` A)" and "Lcm (f ` A)" are printed with optional
|
haftmann@62358
|
627 |
comprehension-like syntax analogously to "Inf (f ` A)" and "Sup (f ` A)".
|
haftmann@62358
|
628 |
|
haftmann@62345
|
629 |
* Class semiring_Lcd merged into semiring_Gcd. INCOMPATIBILITY.
|
haftmann@62345
|
630 |
|
hoelzl@62376
|
631 |
* The type class ordered_comm_monoid_add is now called
|
hoelzl@62376
|
632 |
ordered_cancel_comm_monoid_add. A new type class ordered_comm_monoid_add is
|
hoelzl@62376
|
633 |
introduced as the combination of ordered_ab_semigroup_add + comm_monoid_add.
|
hoelzl@62376
|
634 |
INCOMPATIBILITY.
|
hoelzl@62376
|
635 |
|
hoelzl@62376
|
636 |
* Introduced the type classes canonically_ordered_comm_monoid_add and dioid.
|
hoelzl@62376
|
637 |
|
Mathias@63456
|
638 |
* Introduced the type class ordered_ab_semigroup_monoid_add_imp_le. When
|
Mathias@63456
|
639 |
instantiating linordered_semiring_strict and ordered_ab_group_add, an explicit
|
Mathias@63456
|
640 |
instantiation of ordered_ab_semigroup_monoid_add_imp_le might be
|
Mathias@63456
|
641 |
required.
|
Mathias@63456
|
642 |
INCOMPATIBILITY.
|
Mathias@63456
|
643 |
|
hoelzl@62376
|
644 |
* Added topological_monoid
|
hoelzl@62376
|
645 |
|
Andreas@62649
|
646 |
* Library/Complete_Partial_Order2.thy provides reasoning support for
|
Andreas@62649
|
647 |
proofs about monotonicity and continuity in chain-complete partial
|
Andreas@62649
|
648 |
orders and about admissibility conditions for fixpoint inductions.
|
Andreas@62649
|
649 |
|
haftmann@62352
|
650 |
* Library/Polynomial.thy contains also derivation of polynomials
|
haftmann@62352
|
651 |
but not gcd/lcm on polynomials over fields. This has been moved
|
haftmann@62352
|
652 |
to a separate theory Library/Polynomial_GCD_euclidean.thy, to
|
haftmann@62352
|
653 |
pave way for a possible future different type class instantiation
|
haftmann@62352
|
654 |
for polynomials over factorial rings. INCOMPATIBILITY.
|
haftmann@62352
|
655 |
|
nipkow@63155
|
656 |
* Library/Sublist.thy: added function "prefixes" and renamed
|
nipkow@63173
|
657 |
prefixeq -> prefix
|
nipkow@63173
|
658 |
prefix -> strict_prefix
|
nipkow@63173
|
659 |
suffixeq -> suffix
|
nipkow@63173
|
660 |
suffix -> strict_suffix
|
nipkow@63173
|
661 |
Added theory of longest common prefixes.
|
nipkow@63117
|
662 |
|
haftmann@62348
|
663 |
* Dropped various legacy fact bindings, whose replacements are often
|
haftmann@62348
|
664 |
of a more general type also:
|
haftmann@62348
|
665 |
lcm_left_commute_nat ~> lcm.left_commute
|
haftmann@62348
|
666 |
lcm_left_commute_int ~> lcm.left_commute
|
haftmann@62348
|
667 |
gcd_left_commute_nat ~> gcd.left_commute
|
haftmann@62348
|
668 |
gcd_left_commute_int ~> gcd.left_commute
|
haftmann@62348
|
669 |
gcd_greatest_iff_nat ~> gcd_greatest_iff
|
haftmann@62348
|
670 |
gcd_greatest_iff_int ~> gcd_greatest_iff
|
haftmann@62348
|
671 |
coprime_dvd_mult_nat ~> coprime_dvd_mult
|
haftmann@62348
|
672 |
coprime_dvd_mult_int ~> coprime_dvd_mult
|
haftmann@62348
|
673 |
zpower_numeral_even ~> power_numeral_even
|
haftmann@62348
|
674 |
gcd_mult_cancel_nat ~> gcd_mult_cancel
|
haftmann@62348
|
675 |
gcd_mult_cancel_int ~> gcd_mult_cancel
|
haftmann@62348
|
676 |
div_gcd_coprime_nat ~> div_gcd_coprime
|
haftmann@62348
|
677 |
div_gcd_coprime_int ~> div_gcd_coprime
|
haftmann@62348
|
678 |
zpower_numeral_odd ~> power_numeral_odd
|
haftmann@62348
|
679 |
zero_less_int_conv ~> of_nat_0_less_iff
|
haftmann@62348
|
680 |
gcd_greatest_nat ~> gcd_greatest
|
haftmann@62348
|
681 |
gcd_greatest_int ~> gcd_greatest
|
haftmann@62348
|
682 |
coprime_mult_nat ~> coprime_mult
|
haftmann@62348
|
683 |
coprime_mult_int ~> coprime_mult
|
haftmann@62348
|
684 |
lcm_commute_nat ~> lcm.commute
|
haftmann@62348
|
685 |
lcm_commute_int ~> lcm.commute
|
haftmann@62348
|
686 |
int_less_0_conv ~> of_nat_less_0_iff
|
haftmann@62348
|
687 |
gcd_commute_nat ~> gcd.commute
|
haftmann@62348
|
688 |
gcd_commute_int ~> gcd.commute
|
haftmann@62348
|
689 |
Gcd_insert_nat ~> Gcd_insert
|
haftmann@62348
|
690 |
Gcd_insert_int ~> Gcd_insert
|
haftmann@62348
|
691 |
of_int_int_eq ~> of_int_of_nat_eq
|
haftmann@62348
|
692 |
lcm_least_nat ~> lcm_least
|
haftmann@62348
|
693 |
lcm_least_int ~> lcm_least
|
haftmann@62348
|
694 |
lcm_assoc_nat ~> lcm.assoc
|
haftmann@62348
|
695 |
lcm_assoc_int ~> lcm.assoc
|
haftmann@62348
|
696 |
int_le_0_conv ~> of_nat_le_0_iff
|
haftmann@62348
|
697 |
int_eq_0_conv ~> of_nat_eq_0_iff
|
haftmann@62348
|
698 |
Gcd_empty_nat ~> Gcd_empty
|
haftmann@62348
|
699 |
Gcd_empty_int ~> Gcd_empty
|
haftmann@62348
|
700 |
gcd_assoc_nat ~> gcd.assoc
|
haftmann@62348
|
701 |
gcd_assoc_int ~> gcd.assoc
|
haftmann@62348
|
702 |
zero_zle_int ~> of_nat_0_le_iff
|
haftmann@62348
|
703 |
lcm_dvd2_nat ~> dvd_lcm2
|
haftmann@62348
|
704 |
lcm_dvd2_int ~> dvd_lcm2
|
haftmann@62348
|
705 |
lcm_dvd1_nat ~> dvd_lcm1
|
haftmann@62348
|
706 |
lcm_dvd1_int ~> dvd_lcm1
|
haftmann@62348
|
707 |
gcd_zero_nat ~> gcd_eq_0_iff
|
haftmann@62348
|
708 |
gcd_zero_int ~> gcd_eq_0_iff
|
haftmann@62348
|
709 |
gcd_dvd2_nat ~> gcd_dvd2
|
haftmann@62348
|
710 |
gcd_dvd2_int ~> gcd_dvd2
|
haftmann@62348
|
711 |
gcd_dvd1_nat ~> gcd_dvd1
|
haftmann@62348
|
712 |
gcd_dvd1_int ~> gcd_dvd1
|
haftmann@62348
|
713 |
int_numeral ~> of_nat_numeral
|
haftmann@62348
|
714 |
lcm_ac_nat ~> ac_simps
|
haftmann@62348
|
715 |
lcm_ac_int ~> ac_simps
|
haftmann@62348
|
716 |
gcd_ac_nat ~> ac_simps
|
haftmann@62348
|
717 |
gcd_ac_int ~> ac_simps
|
haftmann@62348
|
718 |
abs_int_eq ~> abs_of_nat
|
haftmann@62348
|
719 |
zless_int ~> of_nat_less_iff
|
haftmann@62348
|
720 |
zdiff_int ~> of_nat_diff
|
haftmann@62348
|
721 |
zadd_int ~> of_nat_add
|
haftmann@62348
|
722 |
int_mult ~> of_nat_mult
|
haftmann@62348
|
723 |
int_Suc ~> of_nat_Suc
|
haftmann@62348
|
724 |
inj_int ~> inj_of_nat
|
haftmann@62348
|
725 |
int_1 ~> of_nat_1
|
haftmann@62348
|
726 |
int_0 ~> of_nat_0
|
haftmann@62353
|
727 |
Lcm_empty_nat ~> Lcm_empty
|
haftmann@62353
|
728 |
Lcm_empty_int ~> Lcm_empty
|
haftmann@62353
|
729 |
Lcm_insert_nat ~> Lcm_insert
|
haftmann@62353
|
730 |
Lcm_insert_int ~> Lcm_insert
|
haftmann@62353
|
731 |
comp_fun_idem_gcd_nat ~> comp_fun_idem_gcd
|
haftmann@62353
|
732 |
comp_fun_idem_gcd_int ~> comp_fun_idem_gcd
|
haftmann@62353
|
733 |
comp_fun_idem_lcm_nat ~> comp_fun_idem_lcm
|
haftmann@62353
|
734 |
comp_fun_idem_lcm_int ~> comp_fun_idem_lcm
|
haftmann@62353
|
735 |
Lcm_eq_0 ~> Lcm_eq_0_I
|
haftmann@62353
|
736 |
Lcm0_iff ~> Lcm_0_iff
|
haftmann@62353
|
737 |
Lcm_dvd_int ~> Lcm_least
|
haftmann@62353
|
738 |
divides_mult_nat ~> divides_mult
|
haftmann@62353
|
739 |
divides_mult_int ~> divides_mult
|
haftmann@62353
|
740 |
lcm_0_nat ~> lcm_0_right
|
haftmann@62353
|
741 |
lcm_0_int ~> lcm_0_right
|
haftmann@62353
|
742 |
lcm_0_left_nat ~> lcm_0_left
|
haftmann@62353
|
743 |
lcm_0_left_int ~> lcm_0_left
|
haftmann@62353
|
744 |
dvd_gcd_D1_nat ~> dvd_gcdD1
|
haftmann@62353
|
745 |
dvd_gcd_D1_int ~> dvd_gcdD1
|
haftmann@62353
|
746 |
dvd_gcd_D2_nat ~> dvd_gcdD2
|
haftmann@62353
|
747 |
dvd_gcd_D2_int ~> dvd_gcdD2
|
haftmann@62353
|
748 |
coprime_dvd_mult_iff_nat ~> coprime_dvd_mult_iff
|
haftmann@62353
|
749 |
coprime_dvd_mult_iff_int ~> coprime_dvd_mult_iff
|
haftmann@62348
|
750 |
realpow_minus_mult ~> power_minus_mult
|
haftmann@62348
|
751 |
realpow_Suc_le_self ~> power_Suc_le_self
|
haftmann@62353
|
752 |
dvd_Gcd, dvd_Gcd_nat, dvd_Gcd_int removed in favour of Gcd_greatest
|
haftmann@62347
|
753 |
INCOMPATIBILITY.
|
haftmann@62347
|
754 |
|
wenzelm@62479
|
755 |
* Session HOL-NSA has been renamed to HOL-Nonstandard_Analysis.
|
wenzelm@62479
|
756 |
|
hoelzl@62975
|
757 |
* In HOL-Probability the type of emeasure and nn_integral was changed
|
hoelzl@62975
|
758 |
from ereal to ennreal:
|
hoelzl@62975
|
759 |
emeasure :: 'a measure => 'a set => ennreal
|
hoelzl@62975
|
760 |
nn_integral :: 'a measure => ('a => ennreal) => ennreal
|
hoelzl@62976
|
761 |
INCOMPATIBILITY.
|
blanchet@62327
|
762 |
|
wenzelm@63198
|
763 |
|
wenzelm@62498
|
764 |
*** ML ***
|
wenzelm@62498
|
765 |
|
wenzelm@63669
|
766 |
* ML antiquotation @{path} is superseded by @{file}, which ensures that
|
wenzelm@63669
|
767 |
the argument is a plain file. Minor INCOMPATIBILITY.
|
wenzelm@63669
|
768 |
|
wenzelm@63227
|
769 |
* Integer.gcd and Integer.lcm use efficient operations from the Poly/ML
|
wenzelm@63228
|
770 |
library (notably for big integers). Subtle change of semantics:
|
wenzelm@63228
|
771 |
Integer.gcd and Integer.lcm both normalize the sign, results are never
|
wenzelm@63228
|
772 |
negative. This coincides with the definitions in HOL/GCD.thy.
|
wenzelm@63228
|
773 |
INCOMPATIBILITY.
|
wenzelm@63227
|
774 |
|
wenzelm@63212
|
775 |
* Structure Rat for rational numbers is now an integral part of
|
wenzelm@63215
|
776 |
Isabelle/ML, with special notation @int/nat or @int for numerals (an
|
wenzelm@63215
|
777 |
abbreviation for antiquotation @{Pure.rat argument}) and ML pretty
|
wenzelm@63212
|
778 |
printing. Standard operations on type Rat.rat are provided via ad-hoc
|
wenzelm@63215
|
779 |
overloading of + - * / < <= > >= ~ abs. INCOMPATIBILITY, need to
|
wenzelm@63212
|
780 |
use + instead of +/ etc. Moreover, exception Rat.DIVZERO has been
|
wenzelm@63212
|
781 |
superseded by General.Div.
|
wenzelm@63198
|
782 |
|
wenzelm@62861
|
783 |
* The ML function "ML" provides easy access to run-time compilation.
|
wenzelm@62861
|
784 |
This is particularly useful for conditional compilation, without
|
wenzelm@62861
|
785 |
requiring separate files.
|
wenzelm@62861
|
786 |
|
wenzelm@62851
|
787 |
* Low-level ML system structures (like PolyML and RunCall) are no longer
|
wenzelm@62886
|
788 |
exposed to Isabelle/ML user-space. Potential INCOMPATIBILITY.
|
wenzelm@62851
|
789 |
|
wenzelm@62662
|
790 |
* Antiquotation @{make_string} is available during Pure bootstrap --
|
wenzelm@62662
|
791 |
with approximative output quality.
|
wenzelm@62662
|
792 |
|
wenzelm@62498
|
793 |
* Option ML_exception_debugger controls detailed exception trace via the
|
wenzelm@62498
|
794 |
Poly/ML debugger. Relevant ML modules need to be compiled beforehand
|
wenzelm@62498
|
795 |
with ML_file_debug, or with ML_file and option ML_debugger enabled. Note
|
wenzelm@62498
|
796 |
debugger information requires consirable time and space: main
|
wenzelm@62498
|
797 |
Isabelle/HOL with full debugger support may need ML_system_64.
|
wenzelm@62498
|
798 |
|
wenzelm@62514
|
799 |
* Local_Theory.restore has been renamed to Local_Theory.reset to
|
wenzelm@62514
|
800 |
emphasize its disruptive impact on the cumulative context, notably the
|
wenzelm@62514
|
801 |
scope of 'private' or 'qualified' names. Note that Local_Theory.reset is
|
wenzelm@62514
|
802 |
only appropriate when targets are managed, e.g. starting from a global
|
wenzelm@62514
|
803 |
theory and returning to it. Regular definitional packages should use
|
wenzelm@62514
|
804 |
balanced blocks of Local_Theory.open_target versus
|
wenzelm@62514
|
805 |
Local_Theory.close_target instead. Rare INCOMPATIBILITY.
|
wenzelm@62514
|
806 |
|
wenzelm@62519
|
807 |
* Structure TimeLimit (originally from the SML/NJ library) has been
|
wenzelm@62519
|
808 |
replaced by structure Timeout, with slightly different signature.
|
wenzelm@62519
|
809 |
INCOMPATIBILITY.
|
wenzelm@62519
|
810 |
|
wenzelm@62551
|
811 |
* Discontinued cd and pwd operations, which are not well-defined in a
|
wenzelm@62551
|
812 |
multi-threaded environment. Note that files are usually located
|
wenzelm@62551
|
813 |
relatively to the master directory of a theory (see also
|
wenzelm@62551
|
814 |
File.full_path). Potential INCOMPATIBILITY.
|
wenzelm@62551
|
815 |
|
wenzelm@63352
|
816 |
* Binding.empty_atts supersedes Thm.empty_binding and
|
wenzelm@63352
|
817 |
Attrib.empty_binding. Minor INCOMPATIBILITY.
|
wenzelm@63352
|
818 |
|
wenzelm@62498
|
819 |
|
wenzelm@62354
|
820 |
*** System ***
|
wenzelm@62354
|
821 |
|
wenzelm@62840
|
822 |
* Many Isabelle tools that require a Java runtime system refer to the
|
wenzelm@62840
|
823 |
settings ISABELLE_TOOL_JAVA_OPTIONS32 / ISABELLE_TOOL_JAVA_OPTIONS64,
|
wenzelm@62840
|
824 |
depending on the underlying platform. The settings for "isabelle build"
|
wenzelm@62840
|
825 |
ISABELLE_BUILD_JAVA_OPTIONS32 / ISABELLE_BUILD_JAVA_OPTIONS64 have been
|
wenzelm@62840
|
826 |
discontinued. Potential INCOMPATIBILITY.
|
wenzelm@62840
|
827 |
|
wenzelm@62591
|
828 |
* The Isabelle system environment always ensures that the main
|
wenzelm@62591
|
829 |
executables are found within the shell search $PATH: "isabelle" and
|
wenzelm@62591
|
830 |
"isabelle_scala_script".
|
wenzelm@62591
|
831 |
|
wenzelm@63226
|
832 |
* Isabelle tools may consist of .scala files: the Scala compiler is
|
wenzelm@63226
|
833 |
invoked on the spot. The source needs to define some object that extends
|
wenzelm@63226
|
834 |
Isabelle_Tool.Body.
|
wenzelm@63226
|
835 |
|
wenzelm@62591
|
836 |
* The Isabelle ML process is now managed directly by Isabelle/Scala, and
|
wenzelm@62591
|
837 |
shell scripts merely provide optional command-line access. In
|
wenzelm@62591
|
838 |
particular:
|
wenzelm@62591
|
839 |
|
wenzelm@62591
|
840 |
. Scala module ML_Process to connect to the raw ML process,
|
wenzelm@62591
|
841 |
with interaction via stdin/stdout/stderr or in batch mode;
|
wenzelm@62591
|
842 |
. command-line tool "isabelle console" as interactive wrapper;
|
wenzelm@62591
|
843 |
. command-line tool "isabelle process" as batch mode wrapper.
|
wenzelm@62588
|
844 |
|
wenzelm@62588
|
845 |
* The executable "isabelle_process" has been discontinued. Tools and
|
wenzelm@62588
|
846 |
prover front-ends should use ML_Process or Isabelle_Process in
|
wenzelm@62591
|
847 |
Isabelle/Scala. INCOMPATIBILITY.
|
wenzelm@62588
|
848 |
|
wenzelm@62588
|
849 |
* New command-line tool "isabelle process" supports ML evaluation of
|
wenzelm@62588
|
850 |
literal expressions (option -e) or files (option -f) in the context of a
|
wenzelm@62588
|
851 |
given heap image. Errors lead to premature exit of the ML process with
|
wenzelm@62588
|
852 |
return code 1.
|
wenzelm@62588
|
853 |
|
wenzelm@62588
|
854 |
* Command-line tool "isabelle console" provides option -r to help to
|
wenzelm@62588
|
855 |
bootstrapping Isabelle/Pure interactively.
|
wenzelm@62588
|
856 |
|
wenzelm@62588
|
857 |
* Command-line tool "isabelle yxml" has been discontinued.
|
wenzelm@62588
|
858 |
INCOMPATIBILITY, use operations from the modules "XML" and "YXML" in
|
wenzelm@62588
|
859 |
Isabelle/ML or Isabelle/Scala.
|
wenzelm@62588
|
860 |
|
wenzelm@62549
|
861 |
* File.bash_string, File.bash_path etc. represent Isabelle/ML and
|
wenzelm@62549
|
862 |
Isabelle/Scala strings authentically within GNU bash. This is useful to
|
wenzelm@62549
|
863 |
produce robust shell scripts under program control, without worrying
|
wenzelm@62549
|
864 |
about spaces or special characters. Note that user output works via
|
wenzelm@62549
|
865 |
Path.print (ML) or Path.toString (Scala). INCOMPATIBILITY, the old (and
|
wenzelm@62549
|
866 |
less versatile) operations File.shell_quote, File.shell_path etc. have
|
wenzelm@62549
|
867 |
been discontinued.
|
wenzelm@62549
|
868 |
|
wenzelm@62591
|
869 |
* SML/NJ and old versions of Poly/ML are no longer supported.
|
wenzelm@62591
|
870 |
|
wenzelm@62642
|
871 |
* Poly/ML heaps now follow the hierarchy of sessions, and thus require
|
wenzelm@62642
|
872 |
much less disk space.
|
wenzelm@62642
|
873 |
|
wenzelm@63827
|
874 |
* System option "checkpoint" helps to fine-tune the global heap space
|
wenzelm@63827
|
875 |
management of isabelle build. This is relevant for big sessions that may
|
wenzelm@63827
|
876 |
exhaust the small 32-bit address space of the ML process (which is used
|
wenzelm@63827
|
877 |
by default).
|
wenzelm@63827
|
878 |
|
wenzelm@62354
|
879 |
|
wenzelm@62354
|
880 |
|
wenzelm@62031
|
881 |
New in Isabelle2016 (February 2016)
|
wenzelm@62016
|
882 |
-----------------------------------
|
wenzelm@60138
|
883 |
|
wenzelm@61337
|
884 |
*** General ***
|
wenzelm@61337
|
885 |
|
wenzelm@62168
|
886 |
* Eisbach is now based on Pure instead of HOL. Objects-logics may import
|
wenzelm@62168
|
887 |
either the theory ~~/src/HOL/Eisbach/Eisbach (for HOL etc.) or
|
wenzelm@62168
|
888 |
~~/src/HOL/Eisbach/Eisbach_Old_Appl_Syntax (for FOL, ZF etc.). Note that
|
wenzelm@62168
|
889 |
the HOL-Eisbach session located in ~~/src/HOL/Eisbach/ contains further
|
wenzelm@62168
|
890 |
examples that do require HOL.
|
wenzelm@62168
|
891 |
|
wenzelm@62157
|
892 |
* Better resource usage on all platforms (Linux, Windows, Mac OS X) for
|
wenzelm@62157
|
893 |
both Isabelle/ML and Isabelle/Scala. Slightly reduced heap space usage.
|
wenzelm@62157
|
894 |
|
wenzelm@62017
|
895 |
* Former "xsymbols" syntax with Isabelle symbols is used by default,
|
wenzelm@62017
|
896 |
without any special print mode. Important ASCII replacement syntax
|
wenzelm@62017
|
897 |
remains available under print mode "ASCII", but less important syntax
|
wenzelm@62017
|
898 |
has been removed (see below).
|
wenzelm@62017
|
899 |
|
wenzelm@62109
|
900 |
* Support for more arrow symbols, with rendering in LaTeX and Isabelle
|
wenzelm@62109
|
901 |
fonts: \<Lleftarrow> \<Rrightarrow> \<longlongleftarrow> \<longlongrightarrow> \<longlonglongleftarrow> \<longlonglongrightarrow>.
|
wenzelm@62017
|
902 |
|
wenzelm@62108
|
903 |
* Special notation \<struct> for the first implicit 'structure' in the
|
wenzelm@62108
|
904 |
context has been discontinued. Rare INCOMPATIBILITY, use explicit
|
wenzelm@62108
|
905 |
structure name instead, notably in indexed notation with block-subscript
|
wenzelm@62108
|
906 |
(e.g. \<odot>\<^bsub>A\<^esub>).
|
wenzelm@62108
|
907 |
|
wenzelm@62108
|
908 |
* The glyph for \<diamond> in the IsabelleText font now corresponds better to its
|
wenzelm@62108
|
909 |
counterpart \<box> as quantifier-like symbol. A small diamond is available as
|
wenzelm@62108
|
910 |
\<diamondop>; the old symbol \<struct> loses this rendering and any special
|
wenzelm@62108
|
911 |
meaning.
|
wenzelm@62108
|
912 |
|
wenzelm@62017
|
913 |
* Syntax for formal comments "-- text" now also supports the symbolic
|
wenzelm@62017
|
914 |
form "\<comment> text". Command-line tool "isabelle update_cartouches -c" helps
|
wenzelm@62017
|
915 |
to update old sources.
|
wenzelm@62017
|
916 |
|
wenzelm@61337
|
917 |
* Toplevel theorem statements have been simplified as follows:
|
wenzelm@61337
|
918 |
|
wenzelm@61337
|
919 |
theorems ~> lemmas
|
wenzelm@61337
|
920 |
schematic_lemma ~> schematic_goal
|
wenzelm@61337
|
921 |
schematic_theorem ~> schematic_goal
|
wenzelm@61337
|
922 |
schematic_corollary ~> schematic_goal
|
wenzelm@61337
|
923 |
|
wenzelm@61337
|
924 |
Command-line tool "isabelle update_theorems" updates theory sources
|
wenzelm@61337
|
925 |
accordingly.
|
wenzelm@61337
|
926 |
|
wenzelm@61338
|
927 |
* Toplevel theorem statement 'proposition' is another alias for
|
wenzelm@61338
|
928 |
'theorem'.
|
wenzelm@61338
|
929 |
|
wenzelm@62169
|
930 |
* The old 'defs' command has been removed (legacy since Isabelle2014).
|
wenzelm@62169
|
931 |
INCOMPATIBILITY, use regular 'definition' instead. Overloaded and/or
|
wenzelm@62169
|
932 |
deferred definitions require a surrounding 'overloading' block.
|
wenzelm@62169
|
933 |
|
wenzelm@61337
|
934 |
|
wenzelm@60610
|
935 |
*** Prover IDE -- Isabelle/Scala/jEdit ***
|
wenzelm@60610
|
936 |
|
wenzelm@60986
|
937 |
* IDE support for the source-level debugger of Poly/ML, to work with
|
wenzelm@62253
|
938 |
Isabelle/ML and official Standard ML. Option "ML_debugger" and commands
|
wenzelm@62253
|
939 |
'ML_file_debug', 'ML_file_no_debug', 'SML_file_debug',
|
wenzelm@62253
|
940 |
'SML_file_no_debug' control compilation of sources with or without
|
wenzelm@62253
|
941 |
debugging information. The Debugger panel allows to set breakpoints (via
|
wenzelm@62253
|
942 |
context menu), step through stopped threads, evaluate local ML
|
wenzelm@62253
|
943 |
expressions etc. At least one Debugger view needs to be active to have
|
wenzelm@62253
|
944 |
any effect on the running ML program.
|
wenzelm@60984
|
945 |
|
wenzelm@61804
|
946 |
* The State panel manages explicit proof state output, with dynamic
|
wenzelm@61804
|
947 |
auto-update according to cursor movement. Alternatively, the jEdit
|
wenzelm@61804
|
948 |
action "isabelle.update-state" (shortcut S+ENTER) triggers manual
|
wenzelm@61804
|
949 |
update.
|
wenzelm@61729
|
950 |
|
wenzelm@61729
|
951 |
* The Output panel no longer shows proof state output by default, to
|
wenzelm@61729
|
952 |
avoid GUI overcrowding. INCOMPATIBILITY, use the State panel instead or
|
wenzelm@61729
|
953 |
enable option "editor_output_state".
|
wenzelm@61215
|
954 |
|
wenzelm@61804
|
955 |
* The text overview column (status of errors, warnings etc.) is updated
|
wenzelm@61804
|
956 |
asynchronously, leading to much better editor reactivity. Moreover, the
|
wenzelm@61804
|
957 |
full document node content is taken into account. The width of the
|
wenzelm@61804
|
958 |
column is scaled according to the main text area font, for improved
|
wenzelm@61804
|
959 |
visibility.
|
wenzelm@61804
|
960 |
|
wenzelm@61804
|
961 |
* The main text area no longer changes its color hue in outdated
|
wenzelm@61804
|
962 |
situations. The text overview column takes over the role to indicate
|
wenzelm@61804
|
963 |
unfinished edits in the PIDE pipeline. This avoids flashing text display
|
wenzelm@61804
|
964 |
due to ad-hoc updates by auxiliary GUI components, such as the State
|
wenzelm@61804
|
965 |
panel.
|
wenzelm@61804
|
966 |
|
wenzelm@62254
|
967 |
* Slightly improved scheduling for urgent print tasks (e.g. command
|
wenzelm@62254
|
968 |
state output, interactive queries) wrt. long-running background tasks.
|
wenzelm@62017
|
969 |
|
wenzelm@62017
|
970 |
* Completion of symbols via prefix of \<name> or \<^name> or \name is
|
wenzelm@62017
|
971 |
always possible, independently of the language context. It is never
|
wenzelm@62017
|
972 |
implicit: a popup will show up unconditionally.
|
wenzelm@62017
|
973 |
|
wenzelm@62017
|
974 |
* Additional abbreviations for syntactic completion may be specified in
|
wenzelm@62017
|
975 |
$ISABELLE_HOME/etc/abbrevs and $ISABELLE_HOME_USER/etc/abbrevs, with
|
wenzelm@62017
|
976 |
support for simple templates using ASCII 007 (bell) as placeholder.
|
wenzelm@62017
|
977 |
|
wenzelm@62234
|
978 |
* Symbols \<oplus>, \<Oplus>, \<otimes>, \<Otimes>, \<odot>, \<Odot>, \<ominus>, \<oslash> no longer provide abbreviations for
|
wenzelm@62234
|
979 |
completion like "+o", "*o", ".o" etc. -- due to conflicts with other
|
wenzelm@62234
|
980 |
ASCII syntax. INCOMPATIBILITY, use plain backslash-completion or define
|
wenzelm@62234
|
981 |
suitable abbreviations in $ISABELLE_HOME_USER/etc/abbrevs.
|
wenzelm@62234
|
982 |
|
wenzelm@61483
|
983 |
* Action "isabelle-emph" (with keyboard shortcut C+e LEFT) controls
|
wenzelm@61483
|
984 |
emphasized text style; the effect is visible in document output, not in
|
wenzelm@61483
|
985 |
the editor.
|
wenzelm@61483
|
986 |
|
wenzelm@61483
|
987 |
* Action "isabelle-reset" now uses keyboard shortcut C+e BACK_SPACE,
|
wenzelm@61483
|
988 |
instead of former C+e LEFT.
|
wenzelm@61483
|
989 |
|
wenzelm@61512
|
990 |
* The command-line tool "isabelle jedit" and the isabelle.Main
|
wenzelm@62027
|
991 |
application wrapper treat the default $USER_HOME/Scratch.thy more
|
wenzelm@61512
|
992 |
uniformly, and allow the dummy file argument ":" to open an empty buffer
|
wenzelm@61512
|
993 |
instead.
|
wenzelm@61512
|
994 |
|
wenzelm@62017
|
995 |
* New command-line tool "isabelle jedit_client" allows to connect to an
|
wenzelm@62017
|
996 |
already running Isabelle/jEdit process. This achieves the effect of
|
wenzelm@62017
|
997 |
single-instance applications seen on common GUI desktops.
|
wenzelm@62017
|
998 |
|
wenzelm@61530
|
999 |
* The default look-and-feel for Linux is the traditional "Metal", which
|
wenzelm@61530
|
1000 |
works better with GUI scaling for very high-resolution displays (e.g.
|
wenzelm@61530
|
1001 |
4K). Moreover, it is generally more robust than "Nimbus".
|
wenzelm@61530
|
1002 |
|
wenzelm@62163
|
1003 |
* Update to jedit-5.3.0, with improved GUI scaling and support of
|
wenzelm@62163
|
1004 |
high-resolution displays (e.g. 4K).
|
wenzelm@62163
|
1005 |
|
wenzelm@62034
|
1006 |
* The main Isabelle executable is managed as single-instance Desktop
|
wenzelm@62034
|
1007 |
application uniformly on all platforms: Linux, Windows, Mac OS X.
|
wenzelm@62034
|
1008 |
|
wenzelm@60610
|
1009 |
|
wenzelm@61405
|
1010 |
*** Document preparation ***
|
wenzelm@61405
|
1011 |
|
haftmann@63553
|
1012 |
* Text and ML antiquotation @{locale} for locales, similar to existing
|
haftmann@63553
|
1013 |
antiquotations for classes.
|
haftmann@63553
|
1014 |
|
wenzelm@62017
|
1015 |
* Commands 'paragraph' and 'subparagraph' provide additional section
|
wenzelm@62017
|
1016 |
headings. Thus there are 6 levels of standard headings, as in HTML.
|
wenzelm@62017
|
1017 |
|
wenzelm@62017
|
1018 |
* Command 'text_raw' has been clarified: input text is processed as in
|
wenzelm@62017
|
1019 |
'text' (with antiquotations and control symbols). The key difference is
|
wenzelm@62017
|
1020 |
the lack of the surrounding isabelle markup environment in output.
|
wenzelm@62017
|
1021 |
|
wenzelm@62017
|
1022 |
* Text is structured in paragraphs and nested lists, using notation that
|
wenzelm@62017
|
1023 |
is similar to Markdown. The control symbols for list items are as
|
wenzelm@62017
|
1024 |
follows:
|
wenzelm@62017
|
1025 |
|
wenzelm@62017
|
1026 |
\<^item> itemize
|
wenzelm@62017
|
1027 |
\<^enum> enumerate
|
wenzelm@62017
|
1028 |
\<^descr> description
|
wenzelm@62017
|
1029 |
|
wenzelm@61491
|
1030 |
* There is a new short form for antiquotations with a single argument
|
wenzelm@61491
|
1031 |
that is a cartouche: \<^name>\<open>...\<close> is equivalent to @{name \<open>...\<close>} and
|
wenzelm@61595
|
1032 |
\<open>...\<close> without control symbol is equivalent to @{cartouche \<open>...\<close>}.
|
wenzelm@61595
|
1033 |
\<^name> without following cartouche is equivalent to @{name}. The
|
wenzelm@61501
|
1034 |
standard Isabelle fonts provide glyphs to render important control
|
wenzelm@61501
|
1035 |
symbols, e.g. "\<^verbatim>", "\<^emph>", "\<^bold>".
|
wenzelm@61491
|
1036 |
|
wenzelm@61595
|
1037 |
* Antiquotations @{noindent}, @{smallskip}, @{medskip}, @{bigskip} with
|
wenzelm@61595
|
1038 |
corresponding control symbols \<^noindent>, \<^smallskip>, \<^medskip>, \<^bigskip> specify spacing formally, using
|
wenzelm@61595
|
1039 |
standard LaTeX macros of the same names.
|
wenzelm@61595
|
1040 |
|
wenzelm@61491
|
1041 |
* Antiquotation @{cartouche} in Isabelle/Pure is the same as @{text}.
|
wenzelm@61491
|
1042 |
Consequently, \<open>...\<close> without any decoration prints literal quasi-formal
|
wenzelm@61492
|
1043 |
text. Command-line tool "isabelle update_cartouches -t" helps to update
|
wenzelm@61492
|
1044 |
old sources, by approximative patching of the content of string and
|
wenzelm@61492
|
1045 |
cartouche tokens seen in theory sources.
|
wenzelm@61491
|
1046 |
|
wenzelm@61491
|
1047 |
* The @{text} antiquotation now ignores the antiquotation option
|
wenzelm@61491
|
1048 |
"source". The given text content is output unconditionally, without any
|
wenzelm@61491
|
1049 |
surrounding quotes etc. Subtle INCOMPATIBILITY, put quotes into the
|
wenzelm@61494
|
1050 |
argument where they are really intended, e.g. @{text \<open>"foo"\<close>}. Initial
|
wenzelm@61494
|
1051 |
or terminal spaces are ignored.
|
wenzelm@61491
|
1052 |
|
wenzelm@62017
|
1053 |
* Antiquotations @{emph} and @{bold} output LaTeX source recursively,
|
wenzelm@62017
|
1054 |
adding appropriate text style markup. These may be used in the short
|
wenzelm@62017
|
1055 |
form \<^emph>\<open>...\<close> and \<^bold>\<open>...\<close>.
|
wenzelm@62017
|
1056 |
|
wenzelm@62017
|
1057 |
* Document antiquotation @{footnote} outputs LaTeX source recursively,
|
wenzelm@62017
|
1058 |
marked as \footnote{}. This may be used in the short form \<^footnote>\<open>...\<close>.
|
wenzelm@62017
|
1059 |
|
wenzelm@62017
|
1060 |
* Antiquotation @{verbatim [display]} supports option "indent".
|
wenzelm@62017
|
1061 |
|
wenzelm@62017
|
1062 |
* Antiquotation @{theory_text} prints uninterpreted theory source text
|
wenzelm@62231
|
1063 |
(Isar outer syntax with command keywords etc.). This may be used in the
|
wenzelm@62231
|
1064 |
short form \<^theory_text>\<open>...\<close>. @{theory_text [display]} supports option "indent".
|
wenzelm@62017
|
1065 |
|
wenzelm@62017
|
1066 |
* Antiquotation @{doc ENTRY} provides a reference to the given
|
wenzelm@62017
|
1067 |
documentation, with a hyperlink in the Prover IDE.
|
wenzelm@62017
|
1068 |
|
wenzelm@62017
|
1069 |
* Antiquotations @{command}, @{method}, @{attribute} print checked
|
wenzelm@62017
|
1070 |
entities of the Isar language.
|
wenzelm@62017
|
1071 |
|
wenzelm@61471
|
1072 |
* HTML presentation uses the standard IsabelleText font and Unicode
|
wenzelm@61471
|
1073 |
rendering of Isabelle symbols like Isabelle/Scala/jEdit. The former
|
wenzelm@61488
|
1074 |
print mode "HTML" loses its special meaning.
|
wenzelm@61471
|
1075 |
|
wenzelm@61405
|
1076 |
|
wenzelm@60406
|
1077 |
*** Isar ***
|
wenzelm@60406
|
1078 |
|
wenzelm@62205
|
1079 |
* Local goals ('have', 'show', 'hence', 'thus') allow structured rule
|
wenzelm@62205
|
1080 |
statements like fixes/assumes/shows in theorem specifications, but the
|
wenzelm@62205
|
1081 |
notation is postfix with keywords 'if' (or 'when') and 'for'. For
|
wenzelm@60555
|
1082 |
example:
|
wenzelm@60414
|
1083 |
|
wenzelm@60414
|
1084 |
have result: "C x y"
|
wenzelm@60414
|
1085 |
if "A x" and "B y"
|
wenzelm@60414
|
1086 |
for x :: 'a and y :: 'a
|
wenzelm@60414
|
1087 |
<proof>
|
wenzelm@60414
|
1088 |
|
wenzelm@60449
|
1089 |
The local assumptions are bound to the name "that". The result is
|
wenzelm@60449
|
1090 |
exported from context of the statement as usual. The above roughly
|
wenzelm@60414
|
1091 |
corresponds to a raw proof block like this:
|
wenzelm@60414
|
1092 |
|
wenzelm@60414
|
1093 |
{
|
wenzelm@60414
|
1094 |
fix x :: 'a and y :: 'a
|
wenzelm@60449
|
1095 |
assume that: "A x" "B y"
|
wenzelm@60414
|
1096 |
have "C x y" <proof>
|
wenzelm@60414
|
1097 |
}
|
wenzelm@60414
|
1098 |
note result = this
|
wenzelm@60406
|
1099 |
|
wenzelm@60555
|
1100 |
The keyword 'when' may be used instead of 'if', to indicate 'presume'
|
wenzelm@60555
|
1101 |
instead of 'assume' above.
|
wenzelm@60555
|
1102 |
|
wenzelm@61733
|
1103 |
* Assumptions ('assume', 'presume') allow structured rule statements
|
wenzelm@61733
|
1104 |
using 'if' and 'for', similar to 'have' etc. above. For example:
|
wenzelm@61658
|
1105 |
|
wenzelm@61658
|
1106 |
assume result: "C x y"
|
wenzelm@61658
|
1107 |
if "A x" and "B y"
|
wenzelm@61658
|
1108 |
for x :: 'a and y :: 'a
|
wenzelm@61658
|
1109 |
|
wenzelm@61658
|
1110 |
This assumes "\<And>x y::'a. A x \<Longrightarrow> B y \<Longrightarrow> C x y" and produces a general
|
wenzelm@61658
|
1111 |
result as usual: "A ?x \<Longrightarrow> B ?y \<Longrightarrow> C ?x ?y".
|
wenzelm@61658
|
1112 |
|
wenzelm@61658
|
1113 |
Vacuous quantification in assumptions is omitted, i.e. a for-context
|
wenzelm@61658
|
1114 |
only effects propositions according to actual use of variables. For
|
wenzelm@61658
|
1115 |
example:
|
wenzelm@61658
|
1116 |
|
wenzelm@61658
|
1117 |
assume "A x" and "B y" for x and y
|
wenzelm@61658
|
1118 |
|
wenzelm@61658
|
1119 |
is equivalent to:
|
wenzelm@61658
|
1120 |
|
wenzelm@61658
|
1121 |
assume "\<And>x. A x" and "\<And>y. B y"
|
wenzelm@61658
|
1122 |
|
wenzelm@60595
|
1123 |
* The meaning of 'show' with Pure rule statements has changed: premises
|
wenzelm@60595
|
1124 |
are treated in the sense of 'assume', instead of 'presume'. This means,
|
wenzelm@62205
|
1125 |
a goal like "\<And>x. A x \<Longrightarrow> B x \<Longrightarrow> C x" can be solved completely as
|
wenzelm@62205
|
1126 |
follows:
|
wenzelm@60595
|
1127 |
|
wenzelm@60595
|
1128 |
show "\<And>x. A x \<Longrightarrow> B x \<Longrightarrow> C x"
|
wenzelm@60595
|
1129 |
|
wenzelm@60595
|
1130 |
or:
|
wenzelm@60595
|
1131 |
|
wenzelm@60595
|
1132 |
show "C x" if "A x" "B x" for x
|
wenzelm@60595
|
1133 |
|
wenzelm@60595
|
1134 |
Rare INCOMPATIBILITY, the old behaviour may be recovered as follows:
|
wenzelm@60595
|
1135 |
|
wenzelm@60595
|
1136 |
show "C x" when "A x" "B x" for x
|
wenzelm@60595
|
1137 |
|
wenzelm@60459
|
1138 |
* New command 'consider' states rules for generalized elimination and
|
wenzelm@60459
|
1139 |
case splitting. This is like a toplevel statement "theorem obtains" used
|
wenzelm@60459
|
1140 |
within a proof body; or like a multi-branch 'obtain' without activation
|
wenzelm@60459
|
1141 |
of the local context elements yet.
|
wenzelm@60459
|
1142 |
|
wenzelm@60455
|
1143 |
* Proof method "cases" allows to specify the rule as first entry of
|
wenzelm@60455
|
1144 |
chained facts. This is particularly useful with 'consider':
|
wenzelm@60455
|
1145 |
|
wenzelm@60455
|
1146 |
consider (a) A | (b) B | (c) C <proof>
|
wenzelm@60455
|
1147 |
then have something
|
wenzelm@60455
|
1148 |
proof cases
|
wenzelm@60455
|
1149 |
case a
|
wenzelm@60455
|
1150 |
then show ?thesis <proof>
|
wenzelm@60455
|
1151 |
next
|
wenzelm@60455
|
1152 |
case b
|
wenzelm@60455
|
1153 |
then show ?thesis <proof>
|
wenzelm@60455
|
1154 |
next
|
wenzelm@60455
|
1155 |
case c
|
wenzelm@60455
|
1156 |
then show ?thesis <proof>
|
wenzelm@60455
|
1157 |
qed
|
wenzelm@60455
|
1158 |
|
wenzelm@60565
|
1159 |
* Command 'case' allows fact name and attribute specification like this:
|
wenzelm@60565
|
1160 |
|
wenzelm@60565
|
1161 |
case a: (c xs)
|
wenzelm@60565
|
1162 |
case a [attributes]: (c xs)
|
wenzelm@60565
|
1163 |
|
wenzelm@60565
|
1164 |
Facts that are introduced by invoking the case context are uniformly
|
wenzelm@60565
|
1165 |
qualified by "a"; the same name is used for the cumulative fact. The old
|
wenzelm@60565
|
1166 |
form "case (c xs) [attributes]" is no longer supported. Rare
|
wenzelm@60565
|
1167 |
INCOMPATIBILITY, need to adapt uses of case facts in exotic situations,
|
wenzelm@60565
|
1168 |
and always put attributes in front.
|
wenzelm@60565
|
1169 |
|
wenzelm@60618
|
1170 |
* The standard proof method of commands 'proof' and '..' is now called
|
wenzelm@60618
|
1171 |
"standard" to make semantically clear what it is; the old name "default"
|
wenzelm@60618
|
1172 |
is still available as legacy for some time. Documentation now explains
|
wenzelm@60618
|
1173 |
'..' more accurately as "by standard" instead of "by rule".
|
wenzelm@60618
|
1174 |
|
wenzelm@62017
|
1175 |
* Nesting of Isar goal structure has been clarified: the context after
|
wenzelm@62017
|
1176 |
the initial backwards refinement is retained for the whole proof, within
|
wenzelm@62017
|
1177 |
all its context sections (as indicated via 'next'). This is e.g.
|
wenzelm@62017
|
1178 |
relevant for 'using', 'including', 'supply':
|
wenzelm@62017
|
1179 |
|
wenzelm@62017
|
1180 |
have "A \<and> A" if a: A for A
|
wenzelm@62017
|
1181 |
supply [simp] = a
|
wenzelm@62017
|
1182 |
proof
|
wenzelm@62017
|
1183 |
show A by simp
|
wenzelm@62017
|
1184 |
next
|
wenzelm@62017
|
1185 |
show A by simp
|
wenzelm@62017
|
1186 |
qed
|
wenzelm@62017
|
1187 |
|
wenzelm@62017
|
1188 |
* Command 'obtain' binds term abbreviations (via 'is' patterns) in the
|
wenzelm@62017
|
1189 |
proof body as well, abstracted over relevant parameters.
|
wenzelm@62017
|
1190 |
|
wenzelm@62017
|
1191 |
* Improved type-inference for theorem statement 'obtains': separate
|
wenzelm@62017
|
1192 |
parameter scope for of each clause.
|
wenzelm@62017
|
1193 |
|
wenzelm@62017
|
1194 |
* Term abbreviations via 'is' patterns also work for schematic
|
wenzelm@62017
|
1195 |
statements: result is abstracted over unknowns.
|
wenzelm@62017
|
1196 |
|
wenzelm@60631
|
1197 |
* Command 'subgoal' allows to impose some structure on backward
|
wenzelm@60631
|
1198 |
refinements, to avoid proof scripts degenerating into long of 'apply'
|
wenzelm@60631
|
1199 |
sequences. Further explanations and examples are given in the isar-ref
|
wenzelm@60631
|
1200 |
manual.
|
wenzelm@60631
|
1201 |
|
wenzelm@62017
|
1202 |
* Command 'supply' supports fact definitions during goal refinement
|
wenzelm@62017
|
1203 |
('apply' scripts).
|
wenzelm@62017
|
1204 |
|
wenzelm@61166
|
1205 |
* Proof method "goal_cases" turns the current subgoals into cases within
|
wenzelm@61166
|
1206 |
the context; the conclusion is bound to variable ?case in each case. For
|
wenzelm@61166
|
1207 |
example:
|
wenzelm@60617
|
1208 |
|
wenzelm@60617
|
1209 |
lemma "\<And>x. A x \<Longrightarrow> B x \<Longrightarrow> C x"
|
wenzelm@60622
|
1210 |
and "\<And>y z. U y \<Longrightarrow> V z \<Longrightarrow> W y z"
|
wenzelm@61166
|
1211 |
proof goal_cases
|
wenzelm@60622
|
1212 |
case (1 x)
|
wenzelm@60622
|
1213 |
then show ?case using \<open>A x\<close> \<open>B x\<close> sorry
|
wenzelm@60622
|
1214 |
next
|
wenzelm@60622
|
1215 |
case (2 y z)
|
wenzelm@60622
|
1216 |
then show ?case using \<open>U y\<close> \<open>V z\<close> sorry
|
wenzelm@60622
|
1217 |
qed
|
wenzelm@60622
|
1218 |
|
wenzelm@60622
|
1219 |
lemma "\<And>x. A x \<Longrightarrow> B x \<Longrightarrow> C x"
|
wenzelm@60622
|
1220 |
and "\<And>y z. U y \<Longrightarrow> V z \<Longrightarrow> W y z"
|
wenzelm@61166
|
1221 |
proof goal_cases
|
wenzelm@60617
|
1222 |
case prems: 1
|
wenzelm@60617
|
1223 |
then show ?case using prems sorry
|
wenzelm@60617
|
1224 |
next
|
wenzelm@60617
|
1225 |
case prems: 2
|
wenzelm@60617
|
1226 |
then show ?case using prems sorry
|
wenzelm@60617
|
1227 |
qed
|
wenzelm@60578
|
1228 |
|
wenzelm@60581
|
1229 |
* The undocumented feature of implicit cases goal1, goal2, goal3, etc.
|
wenzelm@60617
|
1230 |
is marked as legacy, and will be removed eventually. The proof method
|
wenzelm@60617
|
1231 |
"goals" achieves a similar effect within regular Isar; often it can be
|
wenzelm@60617
|
1232 |
done more adequately by other means (e.g. 'consider').
|
wenzelm@60581
|
1233 |
|
wenzelm@62017
|
1234 |
* The vacuous fact "TERM x" may be established "by fact" or as `TERM x`
|
wenzelm@62017
|
1235 |
as well, not just "by this" or "." as before.
|
wenzelm@60551
|
1236 |
|
wenzelm@60554
|
1237 |
* Method "sleep" succeeds after a real-time delay (in seconds). This is
|
wenzelm@60554
|
1238 |
occasionally useful for demonstration and testing purposes.
|
wenzelm@60554
|
1239 |
|
wenzelm@60406
|
1240 |
|
wenzelm@60331
|
1241 |
*** Pure ***
|
wenzelm@60331
|
1242 |
|
wenzelm@61606
|
1243 |
* Qualifiers in locale expressions default to mandatory ('!') regardless
|
wenzelm@61606
|
1244 |
of the command. Previously, for 'locale' and 'sublocale' the default was
|
wenzelm@61606
|
1245 |
optional ('?'). The old synatx '!' has been discontinued.
|
wenzelm@61606
|
1246 |
INCOMPATIBILITY, remove '!' and add '?' as required.
|
ballarin@61565
|
1247 |
|
ballarin@61566
|
1248 |
* Keyword 'rewrites' identifies rewrite morphisms in interpretation
|
wenzelm@62017
|
1249 |
commands. Previously, the keyword was 'where'. INCOMPATIBILITY.
|
ballarin@61566
|
1250 |
|
ballarin@61701
|
1251 |
* More gentle suppression of syntax along locale morphisms while
|
wenzelm@62017
|
1252 |
printing terms. Previously 'abbreviation' and 'notation' declarations
|
wenzelm@62017
|
1253 |
would be suppressed for morphisms except term identity. Now
|
ballarin@61701
|
1254 |
'abbreviation' is also kept for morphims that only change the involved
|
wenzelm@62017
|
1255 |
parameters, and only 'notation' is suppressed. This can be of great help
|
wenzelm@62017
|
1256 |
when working with complex locale hierarchies, because proof states are
|
wenzelm@62017
|
1257 |
displayed much more succinctly. It also means that only notation needs
|
wenzelm@62017
|
1258 |
to be redeclared if desired, as illustrated by this example:
|
ballarin@61701
|
1259 |
|
ballarin@61701
|
1260 |
locale struct = fixes composition :: "'a => 'a => 'a" (infixl "\<cdot>" 65)
|
ballarin@61701
|
1261 |
begin
|
ballarin@61701
|
1262 |
definition derived (infixl "\<odot>" 65) where ...
|
ballarin@61701
|
1263 |
end
|
ballarin@61701
|
1264 |
|
ballarin@61701
|
1265 |
locale morphism =
|
ballarin@61701
|
1266 |
left: struct composition + right: struct composition'
|
ballarin@61701
|
1267 |
for composition (infix "\<cdot>" 65) and composition' (infix "\<cdot>''" 65)
|
ballarin@61701
|
1268 |
begin
|
ballarin@61701
|
1269 |
notation right.derived ("\<odot>''")
|
ballarin@61701
|
1270 |
end
|
ballarin@61701
|
1271 |
|
wenzelm@61895
|
1272 |
* Command 'global_interpretation' issues interpretations into global
|
wenzelm@61895
|
1273 |
theories, with optional rewrite definitions following keyword 'defines'.
|
wenzelm@61895
|
1274 |
|
wenzelm@61895
|
1275 |
* Command 'sublocale' accepts optional rewrite definitions after keyword
|
haftmann@61675
|
1276 |
'defines'.
|
haftmann@61675
|
1277 |
|
wenzelm@61895
|
1278 |
* Command 'permanent_interpretation' has been discontinued. Use
|
wenzelm@61895
|
1279 |
'global_interpretation' or 'sublocale' instead. INCOMPATIBILITY.
|
haftmann@61670
|
1280 |
|
wenzelm@61252
|
1281 |
* Command 'print_definitions' prints dependencies of definitional
|
wenzelm@61252
|
1282 |
specifications. This functionality used to be part of 'print_theory'.
|
wenzelm@61252
|
1283 |
|
wenzelm@60331
|
1284 |
* Configuration option rule_insts_schematic has been discontinued
|
wenzelm@62017
|
1285 |
(intermediate legacy feature in Isabelle2015). INCOMPATIBILITY.
|
wenzelm@60331
|
1286 |
|
wenzelm@62205
|
1287 |
* Abbreviations in type classes now carry proper sort constraint. Rare
|
wenzelm@62205
|
1288 |
INCOMPATIBILITY in situations where the previous misbehaviour has been
|
wenzelm@62205
|
1289 |
exploited.
|
haftmann@60349
|
1290 |
|
haftmann@60349
|
1291 |
* Refinement of user-space type system in type classes: pseudo-local
|
wenzelm@62205
|
1292 |
operations behave more similar to abbreviations. Potential
|
haftmann@60349
|
1293 |
INCOMPATIBILITY in exotic situations.
|
haftmann@60349
|
1294 |
|
haftmann@60349
|
1295 |
|
nipkow@60171
|
1296 |
*** HOL ***
|
nipkow@60171
|
1297 |
|
wenzelm@62017
|
1298 |
* The 'typedef' command has been upgraded from a partially checked
|
wenzelm@62017
|
1299 |
"axiomatization", to a full definitional specification that takes the
|
wenzelm@62017
|
1300 |
global collection of overloaded constant / type definitions into
|
wenzelm@62017
|
1301 |
account. Type definitions with open dependencies on overloaded
|
wenzelm@62017
|
1302 |
definitions need to be specified as "typedef (overloaded)". This
|
wenzelm@62017
|
1303 |
provides extra robustness in theory construction. Rare INCOMPATIBILITY.
|
wenzelm@62017
|
1304 |
|
wenzelm@62017
|
1305 |
* Qualification of various formal entities in the libraries is done more
|
wenzelm@62017
|
1306 |
uniformly via "context begin qualified definition ... end" instead of
|
wenzelm@62017
|
1307 |
old-style "hide_const (open) ...". Consequently, both the defined
|
wenzelm@62017
|
1308 |
constant and its defining fact become qualified, e.g. Option.is_none and
|
wenzelm@62017
|
1309 |
Option.is_none_def. Occasional INCOMPATIBILITY in applications.
|
wenzelm@62017
|
1310 |
|
wenzelm@62017
|
1311 |
* Some old and rarely used ASCII replacement syntax has been removed.
|
wenzelm@62017
|
1312 |
INCOMPATIBILITY, standard syntax with symbols should be used instead.
|
wenzelm@62017
|
1313 |
The subsequent commands help to reproduce the old forms, e.g. to
|
wenzelm@62017
|
1314 |
simplify porting old theories:
|
wenzelm@62017
|
1315 |
|
wenzelm@62017
|
1316 |
notation iff (infixr "<->" 25)
|
wenzelm@62017
|
1317 |
|
wenzelm@62017
|
1318 |
notation Times (infixr "<*>" 80)
|
wenzelm@62017
|
1319 |
|
wenzelm@62017
|
1320 |
type_notation Map.map (infixr "~=>" 0)
|
wenzelm@62017
|
1321 |
notation Map.map_comp (infixl "o'_m" 55)
|
wenzelm@62017
|
1322 |
|
wenzelm@62017
|
1323 |
type_notation FinFun.finfun ("(_ =>f /_)" [22, 21] 21)
|
wenzelm@62017
|
1324 |
|
wenzelm@62017
|
1325 |
notation FuncSet.funcset (infixr "->" 60)
|
wenzelm@62017
|
1326 |
notation FuncSet.extensional_funcset (infixr "->\<^sub>E" 60)
|
wenzelm@62017
|
1327 |
|
wenzelm@62017
|
1328 |
notation Omega_Words_Fun.conc (infixr "conc" 65)
|
wenzelm@62017
|
1329 |
|
wenzelm@62017
|
1330 |
notation Preorder.equiv ("op ~~")
|
wenzelm@62017
|
1331 |
and Preorder.equiv ("(_/ ~~ _)" [51, 51] 50)
|
wenzelm@62017
|
1332 |
|
wenzelm@62017
|
1333 |
notation (in topological_space) tendsto (infixr "--->" 55)
|
wenzelm@62017
|
1334 |
notation (in topological_space) LIMSEQ ("((_)/ ----> (_))" [60, 60] 60)
|
wenzelm@62017
|
1335 |
notation LIM ("((_)/ -- (_)/ --> (_))" [60, 0, 60] 60)
|
wenzelm@62017
|
1336 |
|
wenzelm@62017
|
1337 |
notation NSA.approx (infixl "@=" 50)
|
wenzelm@62017
|
1338 |
notation NSLIMSEQ ("((_)/ ----NS> (_))" [60, 60] 60)
|
wenzelm@62017
|
1339 |
notation NSLIM ("((_)/ -- (_)/ --NS> (_))" [60, 0, 60] 60)
|
wenzelm@62017
|
1340 |
|
wenzelm@62017
|
1341 |
* The alternative notation "\<Colon>" for type and sort constraints has been
|
wenzelm@62017
|
1342 |
removed: in LaTeX document output it looks the same as "::".
|
wenzelm@62017
|
1343 |
INCOMPATIBILITY, use plain "::" instead.
|
wenzelm@62017
|
1344 |
|
wenzelm@62017
|
1345 |
* Commands 'inductive' and 'inductive_set' work better when names for
|
wenzelm@62017
|
1346 |
intro rules are omitted: the "cases" and "induct" rules no longer
|
wenzelm@62017
|
1347 |
declare empty case_names, but no case_names at all. This allows to use
|
wenzelm@62017
|
1348 |
numbered cases in proofs, without requiring method "goal_cases".
|
wenzelm@62017
|
1349 |
|
wenzelm@62017
|
1350 |
* Inductive definitions ('inductive', 'coinductive', etc.) expose
|
wenzelm@62017
|
1351 |
low-level facts of the internal construction only if the option
|
wenzelm@62093
|
1352 |
"inductive_internals" is enabled. This refers to the internal predicate
|
wenzelm@62017
|
1353 |
definition and its monotonicity result. Rare INCOMPATIBILITY.
|
wenzelm@62017
|
1354 |
|
wenzelm@62017
|
1355 |
* Recursive function definitions ('fun', 'function', 'partial_function')
|
wenzelm@62017
|
1356 |
expose low-level facts of the internal construction only if the option
|
wenzelm@62205
|
1357 |
"function_internals" is enabled. Its internal inductive definition is
|
wenzelm@62205
|
1358 |
also subject to "inductive_internals". Rare INCOMPATIBILITY.
|
wenzelm@62093
|
1359 |
|
wenzelm@62093
|
1360 |
* BNF datatypes ('datatype', 'codatatype', etc.) expose low-level facts
|
wenzelm@62093
|
1361 |
of the internal construction only if the option "bnf_internals" is
|
wenzelm@62093
|
1362 |
enabled. This supersedes the former option "bnf_note_all". Rare
|
wenzelm@62093
|
1363 |
INCOMPATIBILITY.
|
wenzelm@62017
|
1364 |
|
wenzelm@62017
|
1365 |
* Combinator to represent case distinction on products is named
|
wenzelm@62017
|
1366 |
"case_prod", uniformly, discontinuing any input aliasses. Very popular
|
wenzelm@62017
|
1367 |
theorem aliasses have been retained.
|
wenzelm@62017
|
1368 |
|
haftmann@61424
|
1369 |
Consolidated facts:
|
haftmann@61424
|
1370 |
PairE ~> prod.exhaust
|
haftmann@61424
|
1371 |
Pair_eq ~> prod.inject
|
haftmann@61424
|
1372 |
pair_collapse ~> prod.collapse
|
haftmann@61424
|
1373 |
Pair_fst_snd_eq ~> prod_eq_iff
|
haftmann@61424
|
1374 |
split_twice ~> prod.case_distrib
|
haftmann@61424
|
1375 |
split_weak_cong ~> prod.case_cong_weak
|
haftmann@61424
|
1376 |
split_split ~> prod.split
|
haftmann@61424
|
1377 |
split_split_asm ~> prod.split_asm
|
haftmann@61424
|
1378 |
splitI ~> case_prodI
|
haftmann@61424
|
1379 |
splitD ~> case_prodD
|
haftmann@61424
|
1380 |
splitI2 ~> case_prodI2
|
haftmann@61424
|
1381 |
splitI2' ~> case_prodI2'
|
haftmann@61424
|
1382 |
splitE ~> case_prodE
|
haftmann@61424
|
1383 |
splitE' ~> case_prodE'
|
haftmann@61424
|
1384 |
split_pair ~> case_prod_Pair
|
haftmann@61424
|
1385 |
split_eta ~> case_prod_eta
|
haftmann@61424
|
1386 |
split_comp ~> case_prod_comp
|
haftmann@61424
|
1387 |
mem_splitI ~> mem_case_prodI
|
haftmann@61424
|
1388 |
mem_splitI2 ~> mem_case_prodI2
|
haftmann@61424
|
1389 |
mem_splitE ~> mem_case_prodE
|
haftmann@61424
|
1390 |
The_split ~> The_case_prod
|
haftmann@61424
|
1391 |
cond_split_eta ~> cond_case_prod_eta
|
haftmann@61424
|
1392 |
Collect_split_in_rel_leE ~> Collect_case_prod_in_rel_leE
|
haftmann@61424
|
1393 |
Collect_split_in_rel_leI ~> Collect_case_prod_in_rel_leI
|
haftmann@61424
|
1394 |
in_rel_Collect_split_eq ~> in_rel_Collect_case_prod_eq
|
haftmann@61424
|
1395 |
Collect_split_Grp_eqD ~> Collect_case_prod_Grp_eqD
|
haftmann@61424
|
1396 |
Collect_split_Grp_inD ~> Collect_case_prod_Grp_in
|
haftmann@61424
|
1397 |
Domain_Collect_split ~> Domain_Collect_case_prod
|
haftmann@61424
|
1398 |
Image_Collect_split ~> Image_Collect_case_prod
|
haftmann@61424
|
1399 |
Range_Collect_split ~> Range_Collect_case_prod
|
haftmann@61424
|
1400 |
Eps_split ~> Eps_case_prod
|
haftmann@61424
|
1401 |
Eps_split_eq ~> Eps_case_prod_eq
|
haftmann@61424
|
1402 |
split_rsp ~> case_prod_rsp
|
haftmann@61424
|
1403 |
curry_split ~> curry_case_prod
|
haftmann@61424
|
1404 |
split_curry ~> case_prod_curry
|
wenzelm@62017
|
1405 |
|
haftmann@61424
|
1406 |
Changes in structure HOLogic:
|
haftmann@61424
|
1407 |
split_const ~> case_prod_const
|
haftmann@61424
|
1408 |
mk_split ~> mk_case_prod
|
haftmann@61424
|
1409 |
mk_psplits ~> mk_ptupleabs
|
haftmann@61424
|
1410 |
strip_psplits ~> strip_ptupleabs
|
wenzelm@62017
|
1411 |
|
wenzelm@62017
|
1412 |
INCOMPATIBILITY.
|
wenzelm@62017
|
1413 |
|
wenzelm@62017
|
1414 |
* The coercions to type 'real' have been reorganised. The function
|
wenzelm@62017
|
1415 |
'real' is no longer overloaded, but has type 'nat => real' and
|
wenzelm@62017
|
1416 |
abbreviates of_nat for that type. Also 'real_of_int :: int => real'
|
wenzelm@62017
|
1417 |
abbreviates of_int for that type. Other overloaded instances of 'real'
|
wenzelm@62017
|
1418 |
have been replaced by 'real_of_ereal' and 'real_of_float'.
|
wenzelm@62017
|
1419 |
|
lp15@61694
|
1420 |
Consolidated facts (among others):
|
lp15@61694
|
1421 |
real_of_nat_le_iff -> of_nat_le_iff
|
lp15@61694
|
1422 |
real_of_nat_numeral of_nat_numeral
|
lp15@61694
|
1423 |
real_of_int_zero of_int_0
|
lp15@61694
|
1424 |
real_of_nat_zero of_nat_0
|
lp15@61694
|
1425 |
real_of_one of_int_1
|
lp15@61694
|
1426 |
real_of_int_add of_int_add
|
lp15@61694
|
1427 |
real_of_nat_add of_nat_add
|
lp15@61694
|
1428 |
real_of_int_diff of_int_diff
|
lp15@61694
|
1429 |
real_of_nat_diff of_nat_diff
|
lp15@61694
|
1430 |
floor_subtract floor_diff_of_int
|
lp15@61694
|
1431 |
real_of_int_inject of_int_eq_iff
|
lp15@61694
|
1432 |
real_of_int_gt_zero_cancel_iff of_int_0_less_iff
|
lp15@61694
|
1433 |
real_of_int_ge_zero_cancel_iff of_int_0_le_iff
|
lp15@61694
|
1434 |
real_of_nat_ge_zero of_nat_0_le_iff
|
lp15@61694
|
1435 |
real_of_int_ceiling_ge le_of_int_ceiling
|
lp15@61694
|
1436 |
ceiling_less_eq ceiling_less_iff
|
lp15@61694
|
1437 |
ceiling_le_eq ceiling_le_iff
|
lp15@61694
|
1438 |
less_floor_eq less_floor_iff
|
lp15@61694
|
1439 |
floor_less_eq floor_less_iff
|
lp15@61694
|
1440 |
floor_divide_eq_div floor_divide_of_int_eq
|
lp15@61694
|
1441 |
real_of_int_zero_cancel of_nat_eq_0_iff
|
lp15@61694
|
1442 |
ceiling_real_of_int ceiling_of_int
|
wenzelm@62017
|
1443 |
|
wenzelm@62017
|
1444 |
INCOMPATIBILITY.
|
wenzelm@61143
|
1445 |
|
wenzelm@60841
|
1446 |
* Theory Map: lemma map_of_is_SomeD was a clone of map_of_SomeD and has
|
wenzelm@60841
|
1447 |
been removed. INCOMPATIBILITY.
|
wenzelm@60841
|
1448 |
|
lars@60712
|
1449 |
* Quickcheck setup for finite sets.
|
lars@60712
|
1450 |
|
nipkow@60171
|
1451 |
* Discontinued simp_legacy_precond. Potential INCOMPATIBILITY.
|
wenzelm@60138
|
1452 |
|
blanchet@60306
|
1453 |
* Sledgehammer:
|
blanchet@61318
|
1454 |
- The MaSh relevance filter has been sped up.
|
blanchet@60306
|
1455 |
- Proof reconstruction has been improved, to minimize the incidence of
|
blanchet@60306
|
1456 |
cases where Sledgehammer gives a proof that does not work.
|
blanchet@60306
|
1457 |
- Auto Sledgehammer now minimizes and preplays the results.
|
blanchet@61030
|
1458 |
- Handle Vampire 4.0 proof output without raising exception.
|
blanchet@61043
|
1459 |
- Eliminated "MASH" environment variable. Use the "MaSh" option in
|
blanchet@61043
|
1460 |
Isabelle/jEdit instead. INCOMPATIBILITY.
|
blanchet@61317
|
1461 |
- Eliminated obsolete "blocking" option and related subcommands.
|
blanchet@60306
|
1462 |
|
blanchet@60310
|
1463 |
* Nitpick:
|
blanchet@61325
|
1464 |
- Fixed soundness bug in translation of "finite" predicate.
|
blanchet@61324
|
1465 |
- Fixed soundness bug in "destroy_constrs" optimization.
|
blanchet@62080
|
1466 |
- Fixed soundness bug in translation of "rat" type.
|
blanchet@60310
|
1467 |
- Removed "check_potential" and "check_genuine" options.
|
blanchet@61317
|
1468 |
- Eliminated obsolete "blocking" option.
|
blanchet@60310
|
1469 |
|
wenzelm@62027
|
1470 |
* (Co)datatype package:
|
blanchet@61345
|
1471 |
- New commands "lift_bnf" and "copy_bnf" for lifting (copying) a BNF
|
blanchet@61345
|
1472 |
structure on the raw type to an abstract type defined using typedef.
|
blanchet@61345
|
1473 |
- Always generate "case_transfer" theorem.
|
wenzelm@62235
|
1474 |
- For mutual types, generate slightly stronger "rel_induct",
|
wenzelm@62235
|
1475 |
"rel_coinduct", and "coinduct" theorems. INCOMPATIBLITY.
|
blanchet@61552
|
1476 |
- Allow discriminators and selectors with the same name as the type
|
blanchet@61552
|
1477 |
being defined.
|
blanchet@61552
|
1478 |
- Avoid various internal name clashes (e.g., 'datatype f = f').
|
traytel@60920
|
1479 |
|
wenzelm@62098
|
1480 |
* Transfer: new methods for interactive debugging of 'transfer' and
|
wenzelm@62098
|
1481 |
'transfer_prover': 'transfer_start', 'transfer_step', 'transfer_end',
|
wenzelm@62098
|
1482 |
'transfer_prover_start' and 'transfer_prover_end'.
|
kuncar@61370
|
1483 |
|
kleing@62118
|
1484 |
* New diagnostic command print_record for displaying record definitions.
|
kleing@62118
|
1485 |
|
haftmann@60868
|
1486 |
* Division on integers is bootstrapped directly from division on
|
wenzelm@62017
|
1487 |
naturals and uses generic numeral algorithm for computations. Slight
|
wenzelm@62017
|
1488 |
INCOMPATIBILITY, simproc numeral_divmod replaces and generalizes former
|
wenzelm@62017
|
1489 |
simprocs binary_int_div and binary_int_mod
|
wenzelm@62017
|
1490 |
|
wenzelm@62017
|
1491 |
* Tightened specification of class semiring_no_zero_divisors. Minor
|
haftmann@60516
|
1492 |
INCOMPATIBILITY.
|
haftmann@60516
|
1493 |
|
haftmann@60688
|
1494 |
* Class algebraic_semidom introduces common algebraic notions of
|
wenzelm@62017
|
1495 |
integral (semi)domains, particularly units. Although logically subsumed
|
wenzelm@62017
|
1496 |
by fields, is is not a super class of these in order not to burden
|
wenzelm@62017
|
1497 |
fields with notions that are trivial there.
|
wenzelm@62017
|
1498 |
|
wenzelm@62017
|
1499 |
* Class normalization_semidom specifies canonical representants for
|
wenzelm@62017
|
1500 |
equivalence classes of associated elements in an integral (semi)domain.
|
wenzelm@62017
|
1501 |
This formalizes associated elements as well.
|
haftmann@60688
|
1502 |
|
haftmann@60688
|
1503 |
* Abstract specification of gcd/lcm operations in classes semiring_gcd,
|
wenzelm@62017
|
1504 |
semiring_Gcd, semiring_Lcd. Minor INCOMPATIBILITY: facts gcd_nat.commute
|
wenzelm@62017
|
1505 |
and gcd_int.commute are subsumed by gcd.commute, as well as
|
wenzelm@62017
|
1506 |
gcd_nat.assoc and gcd_int.assoc by gcd.assoc.
|
wenzelm@62017
|
1507 |
|
wenzelm@62017
|
1508 |
* Former constants Fields.divide (_ / _) and Divides.div (_ div _) are
|
wenzelm@62017
|
1509 |
logically unified to Rings.divide in syntactic type class Rings.divide,
|
wenzelm@62017
|
1510 |
with infix syntax (_ div _). Infix syntax (_ / _) for field division is
|
wenzelm@62017
|
1511 |
added later as abbreviation in class Fields.inverse. INCOMPATIBILITY,
|
wenzelm@62017
|
1512 |
instantiations must refer to Rings.divide rather than the former
|
wenzelm@62017
|
1513 |
separate constants, hence infix syntax (_ / _) is usually not available
|
wenzelm@62017
|
1514 |
during instantiation.
|
wenzelm@62017
|
1515 |
|
wenzelm@62017
|
1516 |
* New cancellation simprocs for boolean algebras to cancel complementary
|
wenzelm@62017
|
1517 |
terms for sup and inf. For example, "sup x (sup y (- x))" simplifies to
|
wenzelm@62017
|
1518 |
"top". INCOMPATIBILITY.
|
Andreas@61628
|
1519 |
|
hoelzl@62101
|
1520 |
* Class uniform_space introduces uniform spaces btw topological spaces
|
hoelzl@62101
|
1521 |
and metric spaces. Minor INCOMPATIBILITY: open_<type>_def needs to be
|
wenzelm@62205
|
1522 |
introduced in the form of an uniformity. Some constants are more general
|
wenzelm@62205
|
1523 |
now, it may be necessary to add type class constraints.
|
hoelzl@62101
|
1524 |
|
hoelzl@62101
|
1525 |
open_real_def \<leadsto> open_dist
|
hoelzl@62101
|
1526 |
open_complex_def \<leadsto> open_dist
|
hoelzl@62101
|
1527 |
|
wenzelm@62026
|
1528 |
* Library/Monad_Syntax: notation uses symbols \<bind> and \<then>. INCOMPATIBILITY.
|
wenzelm@62026
|
1529 |
|
Mathias@60397
|
1530 |
* Library/Multiset:
|
Mathias@60397
|
1531 |
- Renamed multiset inclusion operators:
|
Mathias@60397
|
1532 |
< ~> <#
|
blanchet@62208
|
1533 |
> ~> >#
|
Mathias@60397
|
1534 |
<= ~> <=#
|
blanchet@62208
|
1535 |
>= ~> >=#
|
Mathias@60397
|
1536 |
\<le> ~> \<le>#
|
blanchet@62208
|
1537 |
\<ge> ~> \<ge>#
|
Mathias@60397
|
1538 |
INCOMPATIBILITY.
|
blanchet@62209
|
1539 |
- Added multiset inclusion operator syntax:
|
blanchet@62209
|
1540 |
\<subset>#
|
blanchet@62209
|
1541 |
\<subseteq>#
|
blanchet@62209
|
1542 |
\<supset>#
|
blanchet@62209
|
1543 |
\<supseteq>#
|
Mathias@60397
|
1544 |
- "'a multiset" is no longer an instance of the "order",
|
Mathias@60397
|
1545 |
"ordered_ab_semigroup_add_imp_le", "ordered_cancel_comm_monoid_diff",
|
Mathias@60397
|
1546 |
"semilattice_inf", and "semilattice_sup" type classes. The theorems
|
Mathias@60397
|
1547 |
previously provided by these type classes (directly or indirectly)
|
Mathias@60397
|
1548 |
are now available through the "subset_mset" interpretation
|
Mathias@60397
|
1549 |
(e.g. add_mono ~> subset_mset.add_mono).
|
Mathias@60397
|
1550 |
INCOMPATIBILITY.
|
nipkow@60497
|
1551 |
- Renamed conversions:
|
nipkow@60515
|
1552 |
multiset_of ~> mset
|
nipkow@60515
|
1553 |
multiset_of_set ~> mset_set
|
nipkow@60497
|
1554 |
set_of ~> set_mset
|
nipkow@60497
|
1555 |
INCOMPATIBILITY
|
Mathias@60398
|
1556 |
- Renamed lemmas:
|
Mathias@60398
|
1557 |
mset_le_def ~> subseteq_mset_def
|
Mathias@60398
|
1558 |
mset_less_def ~> subset_mset_def
|
Mathias@60400
|
1559 |
less_eq_multiset.rep_eq ~> subseteq_mset_def
|
Mathias@60400
|
1560 |
INCOMPATIBILITY
|
Mathias@60400
|
1561 |
- Removed lemmas generated by lift_definition:
|
wenzelm@62235
|
1562 |
less_eq_multiset.abs_eq, less_eq_multiset.rsp,
|
wenzelm@62235
|
1563 |
less_eq_multiset.transfer, less_eq_multiset_def
|
Mathias@60400
|
1564 |
INCOMPATIBILITY
|
wenzelm@60007
|
1565 |
|
wenzelm@62017
|
1566 |
* Library/Omega_Words_Fun: Infinite words modeled as functions nat \<Rightarrow> 'a.
|
wenzelm@62017
|
1567 |
|
wenzelm@62017
|
1568 |
* Library/Bourbaki_Witt_Fixpoint: Added formalisation of the
|
wenzelm@62017
|
1569 |
Bourbaki-Witt fixpoint theorem for increasing functions in
|
wenzelm@62017
|
1570 |
chain-complete partial orders.
|
wenzelm@62017
|
1571 |
|
wenzelm@62017
|
1572 |
* Library/Old_Recdef: discontinued obsolete 'defer_recdef' command.
|
wenzelm@62017
|
1573 |
Minor INCOMPATIBILITY, use 'function' instead.
|
wenzelm@62017
|
1574 |
|
wenzelm@62065
|
1575 |
* Library/Periodic_Fun: a locale that provides convenient lemmas for
|
wenzelm@62065
|
1576 |
periodic functions.
|
eberlm@62060
|
1577 |
|
wenzelm@62098
|
1578 |
* Library/Formal_Power_Series: proper definition of division (with
|
wenzelm@62098
|
1579 |
remainder) for formal power series; instances for Euclidean Ring and
|
wenzelm@62098
|
1580 |
GCD.
|
eberlm@62086
|
1581 |
|
wenzelm@62084
|
1582 |
* HOL-Imperative_HOL: obsolete theory Legacy_Mrec has been removed.
|
wenzelm@62084
|
1583 |
|
wenzelm@62084
|
1584 |
* HOL-Statespace: command 'statespace' uses mandatory qualifier for
|
wenzelm@62084
|
1585 |
import of parent, as for general 'locale' expressions. INCOMPATIBILITY,
|
wenzelm@62084
|
1586 |
remove '!' and add '?' as required.
|
wenzelm@62084
|
1587 |
|
wenzelm@62237
|
1588 |
* HOL-Decision_Procs: The "approximation" method works with "powr"
|
wenzelm@62237
|
1589 |
(exponentiation on real numbers) again.
|
wenzelm@62237
|
1590 |
|
wenzelm@62084
|
1591 |
* HOL-Multivariate_Analysis: theory Cauchy_Integral_Thm with Contour
|
wenzelm@62084
|
1592 |
integrals (= complex path integrals), Cauchy's integral theorem, winding
|
wenzelm@62084
|
1593 |
numbers and Cauchy's integral formula, Liouville theorem, Fundamental
|
wenzelm@62084
|
1594 |
Theorem of Algebra. Ported from HOL Light.
|
wenzelm@62084
|
1595 |
|
wenzelm@62084
|
1596 |
* HOL-Multivariate_Analysis: topological concepts such as connected
|
wenzelm@62017
|
1597 |
components, homotopic paths and the inside or outside of a set.
|
wenzelm@61121
|
1598 |
|
wenzelm@62084
|
1599 |
* HOL-Multivariate_Analysis: radius of convergence of power series and
|
wenzelm@62065
|
1600 |
various summability tests; Harmonic numbers and the Euler–Mascheroni
|
wenzelm@62065
|
1601 |
constant; the Generalised Binomial Theorem; the complex and real
|
wenzelm@62065
|
1602 |
Gamma/log-Gamma/Digamma/ Polygamma functions and their most important
|
wenzelm@62065
|
1603 |
properties.
|
eberlm@62060
|
1604 |
|
wenzelm@62084
|
1605 |
* HOL-Probability: The central limit theorem based on Levy's uniqueness
|
wenzelm@62084
|
1606 |
and continuity theorems, weak convergence, and characterisitc functions.
|
wenzelm@62084
|
1607 |
|
wenzelm@62084
|
1608 |
* HOL-Data_Structures: new and growing session of standard data
|
wenzelm@62084
|
1609 |
structures.
|
lammich@61178
|
1610 |
|
wenzelm@60479
|
1611 |
|
wenzelm@60793
|
1612 |
*** ML ***
|
wenzelm@60793
|
1613 |
|
wenzelm@62017
|
1614 |
* The following combinators for low-level profiling of the ML runtime
|
wenzelm@62017
|
1615 |
system are available:
|
wenzelm@62017
|
1616 |
|
wenzelm@62017
|
1617 |
profile_time (*CPU time*)
|
wenzelm@62017
|
1618 |
profile_time_thread (*CPU time on this thread*)
|
wenzelm@62017
|
1619 |
profile_allocations (*overall heap allocations*)
|
wenzelm@62017
|
1620 |
|
wenzelm@62017
|
1621 |
* Antiquotation @{undefined} or \<^undefined> inlines (raise Match).
|
wenzelm@62017
|
1622 |
|
wenzelm@62075
|
1623 |
* Antiquotation @{method NAME} inlines the (checked) name of the given
|
wenzelm@62075
|
1624 |
Isar proof method.
|
wenzelm@62075
|
1625 |
|
wenzelm@61922
|
1626 |
* Pretty printing of Poly/ML compiler output in Isabelle has been
|
wenzelm@61922
|
1627 |
improved: proper treatment of break offsets and blocks with consistent
|
wenzelm@61922
|
1628 |
breaks.
|
wenzelm@61922
|
1629 |
|
wenzelm@61268
|
1630 |
* The auxiliary module Pure/display.ML has been eliminated. Its
|
wenzelm@61268
|
1631 |
elementary thm print operations are now in Pure/more_thm.ML and thus
|
wenzelm@61268
|
1632 |
called Thm.pretty_thm, Thm.string_of_thm etc. INCOMPATIBILITY.
|
wenzelm@61268
|
1633 |
|
wenzelm@61144
|
1634 |
* Simproc programming interfaces have been simplified:
|
wenzelm@61144
|
1635 |
Simplifier.make_simproc and Simplifier.define_simproc supersede various
|
wenzelm@61144
|
1636 |
forms of Simplifier.mk_simproc, Simplifier.simproc_global etc. Note that
|
wenzelm@61144
|
1637 |
term patterns for the left-hand sides are specified with implicitly
|
wenzelm@61144
|
1638 |
fixed variables, like top-level theorem statements. INCOMPATIBILITY.
|
wenzelm@61144
|
1639 |
|
wenzelm@60802
|
1640 |
* Instantiation rules have been re-organized as follows:
|
wenzelm@60802
|
1641 |
|
wenzelm@60802
|
1642 |
Thm.instantiate (*low-level instantiation with named arguments*)
|
wenzelm@60802
|
1643 |
Thm.instantiate' (*version with positional arguments*)
|
wenzelm@60802
|
1644 |
|
wenzelm@60802
|
1645 |
Drule.infer_instantiate (*instantiation with type inference*)
|
wenzelm@60802
|
1646 |
Drule.infer_instantiate' (*version with positional arguments*)
|
wenzelm@60802
|
1647 |
|
wenzelm@60802
|
1648 |
The LHS only requires variable specifications, instead of full terms.
|
wenzelm@60802
|
1649 |
Old cterm_instantiate is superseded by infer_instantiate.
|
wenzelm@60802
|
1650 |
INCOMPATIBILITY, need to re-adjust some ML names and types accordingly.
|
wenzelm@60802
|
1651 |
|
wenzelm@60793
|
1652 |
* Old tactic shorthands atac, rtac, etac, dtac, ftac have been
|
wenzelm@60793
|
1653 |
discontinued. INCOMPATIBILITY, use regular assume_tac, resolve_tac etc.
|
wenzelm@60793
|
1654 |
instead (with proper context).
|
wenzelm@60642
|
1655 |
|
wenzelm@60642
|
1656 |
* Thm.instantiate (and derivatives) no longer require the LHS of the
|
wenzelm@60642
|
1657 |
instantiation to be certified: plain variables are given directly.
|
wenzelm@60642
|
1658 |
|
wenzelm@60707
|
1659 |
* Subgoal.SUBPROOF and Subgoal.FOCUS combinators use anonymous
|
wenzelm@60707
|
1660 |
quasi-bound variables (like the Simplifier), instead of accidentally
|
wenzelm@60707
|
1661 |
named local fixes. This has the potential to improve stability of proof
|
wenzelm@60707
|
1662 |
tools, but can also cause INCOMPATIBILITY for tools that don't observe
|
wenzelm@60707
|
1663 |
the proof context discipline.
|
wenzelm@60707
|
1664 |
|
wenzelm@62017
|
1665 |
* Isar proof methods are based on a slightly more general type
|
wenzelm@62017
|
1666 |
context_tactic, which allows to change the proof context dynamically
|
wenzelm@62017
|
1667 |
(e.g. to update cases) and indicate explicit Seq.Error results. Former
|
wenzelm@62017
|
1668 |
METHOD_CASES is superseded by CONTEXT_METHOD; further combinators are
|
wenzelm@62017
|
1669 |
provided in src/Pure/Isar/method.ML for convenience. INCOMPATIBILITY.
|
wenzelm@61887
|
1670 |
|
wenzelm@60642
|
1671 |
|
wenzelm@60983
|
1672 |
*** System ***
|
wenzelm@60983
|
1673 |
|
wenzelm@62525
|
1674 |
* Command-line tool "isabelle console" enables print mode "ASCII".
|
wenzelm@61958
|
1675 |
|
wenzelm@62017
|
1676 |
* Command-line tool "isabelle update_then" expands old Isar command
|
wenzelm@62017
|
1677 |
conflations:
|
wenzelm@62017
|
1678 |
|
wenzelm@62017
|
1679 |
hence ~> then have
|
wenzelm@62017
|
1680 |
thus ~> then show
|
wenzelm@62017
|
1681 |
|
wenzelm@62017
|
1682 |
This syntax is more orthogonal and improves readability and
|
wenzelm@62017
|
1683 |
maintainability of proofs.
|
wenzelm@62017
|
1684 |
|
wenzelm@61602
|
1685 |
* Global session timeout is multiplied by timeout_scale factor. This
|
wenzelm@61602
|
1686 |
allows to adjust large-scale tests (e.g. AFP) to overall hardware
|
wenzelm@61602
|
1687 |
performance.
|
wenzelm@61602
|
1688 |
|
wenzelm@61174
|
1689 |
* Property values in etc/symbols may contain spaces, if written with the
|
wenzelm@62671
|
1690 |
replacement character "␣" (Unicode point 0x2324). For example:
|
wenzelm@62671
|
1691 |
|
wenzelm@62671
|
1692 |
\<star> code: 0x0022c6 group: operator font: Deja␣Vu␣Sans␣Mono
|
wenzelm@61174
|
1693 |
|
wenzelm@60996
|
1694 |
* Java runtime environment for x86_64-windows allows to use larger heap
|
wenzelm@60996
|
1695 |
space.
|
wenzelm@60996
|
1696 |
|
wenzelm@61135
|
1697 |
* Java runtime options are determined separately for 32bit vs. 64bit
|
wenzelm@61135
|
1698 |
platforms as follows.
|
wenzelm@61135
|
1699 |
|
wenzelm@61135
|
1700 |
- Isabelle desktop application: platform-specific files that are
|
wenzelm@61135
|
1701 |
associated with the main app bundle
|
wenzelm@61135
|
1702 |
|
wenzelm@61135
|
1703 |
- isabelle jedit: settings
|
wenzelm@61135
|
1704 |
JEDIT_JAVA_SYSTEM_OPTIONS
|
wenzelm@61135
|
1705 |
JEDIT_JAVA_OPTIONS32 vs. JEDIT_JAVA_OPTIONS64
|
wenzelm@61135
|
1706 |
|
wenzelm@61135
|
1707 |
- isabelle build: settings
|
wenzelm@61135
|
1708 |
ISABELLE_BUILD_JAVA_OPTIONS32 vs. ISABELLE_BUILD_JAVA_OPTIONS64
|
wenzelm@61135
|
1709 |
|
wenzelm@61294
|
1710 |
* Bash shell function "jvmpath" has been renamed to "platform_path": it
|
wenzelm@61294
|
1711 |
is relevant both for Poly/ML and JVM processes.
|
wenzelm@61294
|
1712 |
|
wenzelm@62017
|
1713 |
* Poly/ML default platform architecture may be changed from 32bit to
|
wenzelm@62205
|
1714 |
64bit via system option ML_system_64. A system restart (and rebuild) is
|
wenzelm@62205
|
1715 |
required after change.
|
wenzelm@62017
|
1716 |
|
wenzelm@62017
|
1717 |
* Poly/ML 5.6 runs natively on x86-windows and x86_64-windows, which
|
wenzelm@62017
|
1718 |
both allow larger heap space than former x86-cygwin.
|
wenzelm@62017
|
1719 |
|
wenzelm@62157
|
1720 |
* Heap images are 10-15% smaller due to less wasteful persistent theory
|
wenzelm@62157
|
1721 |
content (using ML type theory_id instead of theory);
|
wenzelm@62157
|
1722 |
|
wenzelm@60983
|
1723 |
|
wenzelm@60479
|
1724 |
|
wenzelm@60010
|
1725 |
New in Isabelle2015 (May 2015)
|
wenzelm@60010
|
1726 |
------------------------------
|
wenzelm@57695
|
1727 |
|
wenzelm@57941
|
1728 |
*** General ***
|
wenzelm@57941
|
1729 |
|
wenzelm@59939
|
1730 |
* Local theory specification commands may have a 'private' or
|
wenzelm@59990
|
1731 |
'qualified' modifier to restrict name space accesses to the local scope,
|
wenzelm@59939
|
1732 |
as provided by some "context begin ... end" block. For example:
|
wenzelm@59926
|
1733 |
|
wenzelm@59926
|
1734 |
context
|
wenzelm@59926
|
1735 |
begin
|
wenzelm@59926
|
1736 |
|
wenzelm@59926
|
1737 |
private definition ...
|
wenzelm@59926
|
1738 |
private lemma ...
|
wenzelm@59926
|
1739 |
|
wenzelm@59990
|
1740 |
qualified definition ...
|
wenzelm@59990
|
1741 |
qualified lemma ...
|
wenzelm@59990
|
1742 |
|
wenzelm@59926
|
1743 |
lemma ...
|
wenzelm@59926
|
1744 |
theorem ...
|
wenzelm@59926
|
1745 |
|
wenzelm@59926
|
1746 |
end
|
wenzelm@59926
|
1747 |
|
wenzelm@59901
|
1748 |
* Command 'experiment' opens an anonymous locale context with private
|
wenzelm@59901
|
1749 |
naming policy.
|
wenzelm@59901
|
1750 |
|
wenzelm@59951
|
1751 |
* Command 'notepad' requires proper nesting of begin/end and its proof
|
wenzelm@59951
|
1752 |
structure in the body: 'oops' is no longer supported here. Minor
|
wenzelm@59951
|
1753 |
INCOMPATIBILITY, use 'sorry' instead.
|
wenzelm@59951
|
1754 |
|
wenzelm@59951
|
1755 |
* Command 'named_theorems' declares a dynamic fact within the context,
|
wenzelm@59951
|
1756 |
together with an attribute to maintain the content incrementally. This
|
wenzelm@59951
|
1757 |
supersedes functor Named_Thms in Isabelle/ML, but with a subtle change
|
wenzelm@59951
|
1758 |
of semantics due to external visual order vs. internal reverse order.
|
wenzelm@59951
|
1759 |
|
wenzelm@59951
|
1760 |
* 'find_theorems': search patterns which are abstractions are
|
wenzelm@59951
|
1761 |
schematically expanded before search. Search results match the naive
|
wenzelm@59951
|
1762 |
expectation more closely, particularly wrt. abbreviations.
|
wenzelm@59951
|
1763 |
INCOMPATIBILITY.
|
wenzelm@59648
|
1764 |
|
wenzelm@59569
|
1765 |
* Commands 'method_setup' and 'attribute_setup' now work within a local
|
wenzelm@59569
|
1766 |
theory context.
|
wenzelm@57941
|
1767 |
|
wenzelm@58928
|
1768 |
* Outer syntax commands are managed authentically within the theory
|
wenzelm@59569
|
1769 |
context, without implicit global state. Potential for accidental
|
wenzelm@58928
|
1770 |
INCOMPATIBILITY, make sure that required theories are really imported.
|
wenzelm@58928
|
1771 |
|
wenzelm@60116
|
1772 |
* Historical command-line terminator ";" is no longer accepted (and
|
wenzelm@60116
|
1773 |
already used differently in Isar). Minor INCOMPATIBILITY, use "isabelle
|
wenzelm@60116
|
1774 |
update_semicolons" to remove obsolete semicolons from old theory
|
wenzelm@60116
|
1775 |
sources.
|
wenzelm@60116
|
1776 |
|
wenzelm@59951
|
1777 |
* Structural composition of proof methods (meth1; meth2) in Isar
|
wenzelm@59951
|
1778 |
corresponds to (tac1 THEN_ALL_NEW tac2) in ML.
|
haftmann@59105
|
1779 |
|
wenzelm@60119
|
1780 |
* The Eisbach proof method language allows to define new proof methods
|
wenzelm@60119
|
1781 |
by combining existing ones with their usual syntax. The "match" proof
|
wenzelm@60119
|
1782 |
method provides basic fact/term matching in addition to
|
wenzelm@60119
|
1783 |
premise/conclusion matching through Subgoal.focus, and binds fact names
|
wenzelm@60288
|
1784 |
from matches as well as term patterns within matches. The Isabelle
|
wenzelm@60288
|
1785 |
documentation provides an entry "eisbach" for the Eisbach User Manual.
|
wenzelm@60288
|
1786 |
Sources and various examples are in ~~/src/HOL/Eisbach/.
|
wenzelm@60119
|
1787 |
|
wenzelm@57941
|
1788 |
|
wenzelm@58524
|
1789 |
*** Prover IDE -- Isabelle/Scala/jEdit ***
|
wenzelm@58524
|
1790 |
|
wenzelm@59569
|
1791 |
* Improved folding mode "isabelle" based on Isar syntax. Alternatively,
|
wenzelm@59569
|
1792 |
the "sidekick" mode may be used for document structure.
|
wenzelm@59569
|
1793 |
|
wenzelm@59569
|
1794 |
* Extended bracket matching based on Isar language structure. System
|
wenzelm@59569
|
1795 |
option jedit_structure_limit determines maximum number of lines to scan
|
wenzelm@59569
|
1796 |
in the buffer.
|
wenzelm@58758
|
1797 |
|
wenzelm@58540
|
1798 |
* Support for BibTeX files: context menu, context-sensitive token
|
wenzelm@58540
|
1799 |
marker, SideKick parser.
|
wenzelm@58524
|
1800 |
|
wenzelm@58551
|
1801 |
* Document antiquotation @{cite} provides formal markup, which is
|
wenzelm@60265
|
1802 |
interpreted semi-formally based on .bib files that happen to be open in
|
wenzelm@60265
|
1803 |
the editor (hyperlinks, completion etc.).
|
wenzelm@58551
|
1804 |
|
wenzelm@58785
|
1805 |
* Less waste of vertical space via negative line spacing (see Global
|
wenzelm@58785
|
1806 |
Options / Text Area).
|
wenzelm@58785
|
1807 |
|
wenzelm@60091
|
1808 |
* Improved graphview panel with optional output of PNG or PDF, for
|
wenzelm@60273
|
1809 |
display of 'thy_deps', 'class_deps' etc.
|
wenzelm@60010
|
1810 |
|
wenzelm@60116
|
1811 |
* The commands 'thy_deps' and 'class_deps' allow optional bounds to
|
wenzelm@60116
|
1812 |
restrict the visualized hierarchy.
|
wenzelm@60095
|
1813 |
|
wenzelm@60072
|
1814 |
* Improved scheduling for asynchronous print commands (e.g. provers
|
wenzelm@60072
|
1815 |
managed by the Sledgehammer panel) wrt. ongoing document processing.
|
wenzelm@60072
|
1816 |
|
wenzelm@58524
|
1817 |
|
wenzelm@59951
|
1818 |
*** Document preparation ***
|
wenzelm@59951
|
1819 |
|
wenzelm@59951
|
1820 |
* Document markup commands 'chapter', 'section', 'subsection',
|
wenzelm@59951
|
1821 |
'subsubsection', 'text', 'txt', 'text_raw' work uniformly in any
|
wenzelm@59951
|
1822 |
context, even before the initial 'theory' command. Obsolete proof
|
wenzelm@59951
|
1823 |
commands 'sect', 'subsect', 'subsubsect', 'txt_raw' have been
|
wenzelm@59951
|
1824 |
discontinued, use 'section', 'subsection', 'subsubsection', 'text_raw'
|
wenzelm@59951
|
1825 |
instead. The old 'header' command is still retained for some time, but
|
wenzelm@59951
|
1826 |
should be replaced by 'chapter', 'section' etc. (using "isabelle
|
wenzelm@59951
|
1827 |
update_header"). Minor INCOMPATIBILITY.
|
wenzelm@59951
|
1828 |
|
wenzelm@60010
|
1829 |
* Official support for "tt" style variants, via \isatt{...} or
|
wenzelm@60010
|
1830 |
\begin{isabellett}...\end{isabellett}. The somewhat fragile \verb or
|
wenzelm@60010
|
1831 |
verbatim environment of LaTeX is no longer used. This allows @{ML} etc.
|
wenzelm@60010
|
1832 |
as argument to other macros (such as footnotes).
|
wenzelm@60010
|
1833 |
|
wenzelm@60010
|
1834 |
* Document antiquotation @{verbatim} prints ASCII text literally in "tt"
|
wenzelm@60010
|
1835 |
style.
|
wenzelm@60010
|
1836 |
|
wenzelm@60010
|
1837 |
* Discontinued obsolete option "document_graph": session_graph.pdf is
|
wenzelm@60010
|
1838 |
produced unconditionally for HTML browser_info and PDF-LaTeX document.
|
wenzelm@60010
|
1839 |
|
wenzelm@59951
|
1840 |
* Diagnostic commands and document markup commands within a proof do not
|
wenzelm@59951
|
1841 |
affect the command tag for output. Thus commands like 'thm' are subject
|
wenzelm@59951
|
1842 |
to proof document structure, and no longer "stick out" accidentally.
|
wenzelm@59951
|
1843 |
Commands 'text' and 'txt' merely differ in the LaTeX style, not their
|
wenzelm@59951
|
1844 |
tags. Potential INCOMPATIBILITY in exotic situations.
|
wenzelm@59951
|
1845 |
|
wenzelm@59951
|
1846 |
* System option "pretty_margin" is superseded by "thy_output_margin",
|
wenzelm@59951
|
1847 |
which is also accessible via document antiquotation option "margin".
|
wenzelm@59951
|
1848 |
Only the margin for document output may be changed, but not the global
|
wenzelm@59951
|
1849 |
pretty printing: that is 76 for plain console output, and adapted
|
wenzelm@59951
|
1850 |
dynamically in GUI front-ends. Implementations of document
|
wenzelm@59951
|
1851 |
antiquotations need to observe the margin explicitly according to
|
wenzelm@59951
|
1852 |
Thy_Output.string_of_margin. Minor INCOMPATIBILITY.
|
wenzelm@59951
|
1853 |
|
wenzelm@60299
|
1854 |
* Specification of 'document_files' in the session ROOT file is
|
wenzelm@60299
|
1855 |
mandatory for document preparation. The legacy mode with implicit
|
wenzelm@60299
|
1856 |
copying of the document/ directory is no longer supported. Minor
|
wenzelm@60299
|
1857 |
INCOMPATIBILITY.
|
wenzelm@60299
|
1858 |
|
wenzelm@59951
|
1859 |
|
haftmann@58202
|
1860 |
*** Pure ***
|
haftmann@58202
|
1861 |
|
wenzelm@59835
|
1862 |
* Proof methods with explicit instantiation ("rule_tac", "subgoal_tac"
|
wenzelm@59835
|
1863 |
etc.) allow an optional context of local variables ('for' declaration):
|
wenzelm@59835
|
1864 |
these variables become schematic in the instantiated theorem; this
|
wenzelm@59835
|
1865 |
behaviour is analogous to 'for' in attributes "where" and "of".
|
wenzelm@59835
|
1866 |
Configuration option rule_insts_schematic (default false) controls use
|
wenzelm@59835
|
1867 |
of schematic variables outside the context. Minor INCOMPATIBILITY,
|
wenzelm@59835
|
1868 |
declare rule_insts_schematic = true temporarily and update to use local
|
wenzelm@59835
|
1869 |
variable declarations or dummy patterns instead.
|
wenzelm@59835
|
1870 |
|
wenzelm@60010
|
1871 |
* Explicit instantiation via attributes "where", "of", and proof methods
|
wenzelm@60010
|
1872 |
"rule_tac" with derivatives like "subgoal_tac" etc. admit dummy patterns
|
wenzelm@60010
|
1873 |
("_") that stand for anonymous local variables.
|
wenzelm@60010
|
1874 |
|
wenzelm@59951
|
1875 |
* Generated schematic variables in standard format of exported facts are
|
wenzelm@59951
|
1876 |
incremented to avoid material in the proof context. Rare
|
wenzelm@59951
|
1877 |
INCOMPATIBILITY, explicit instantiation sometimes needs to refer to
|
wenzelm@59951
|
1878 |
different index.
|
wenzelm@59951
|
1879 |
|
wenzelm@60011
|
1880 |
* Lexical separation of signed and unsigned numerals: categories "num"
|
wenzelm@60011
|
1881 |
and "float" are unsigned. INCOMPATIBILITY: subtle change in precedence
|
wenzelm@60011
|
1882 |
of numeral signs, particularly in expressions involving infix syntax
|
wenzelm@60011
|
1883 |
like "(- 1) ^ n".
|
haftmann@58410
|
1884 |
|
wenzelm@58421
|
1885 |
* Old inner token category "xnum" has been discontinued. Potential
|
wenzelm@58421
|
1886 |
INCOMPATIBILITY for exotic syntax: may use mixfix grammar with "num"
|
wenzelm@58421
|
1887 |
token category instead.
|
wenzelm@58421
|
1888 |
|
haftmann@58202
|
1889 |
|
blanchet@57737
|
1890 |
*** HOL ***
|
blanchet@57737
|
1891 |
|
blanchet@57983
|
1892 |
* New (co)datatype package:
|
blanchet@58374
|
1893 |
- The 'datatype_new' command has been renamed 'datatype'. The old
|
blanchet@58374
|
1894 |
command of that name is now called 'old_datatype' and is provided
|
blanchet@58374
|
1895 |
by "~~/src/HOL/Library/Old_Datatype.thy". See
|
blanchet@58374
|
1896 |
'isabelle doc datatypes' for information on porting.
|
blanchet@58374
|
1897 |
INCOMPATIBILITY.
|
blanchet@57983
|
1898 |
- Renamed theorems:
|
blanchet@57983
|
1899 |
disc_corec ~> corec_disc
|
blanchet@57983
|
1900 |
disc_corec_iff ~> corec_disc_iff
|
blanchet@57983
|
1901 |
disc_exclude ~> distinct_disc
|
blanchet@57983
|
1902 |
disc_exhaust ~> exhaust_disc
|
blanchet@57983
|
1903 |
disc_map_iff ~> map_disc_iff
|
blanchet@57983
|
1904 |
sel_corec ~> corec_sel
|
blanchet@57983
|
1905 |
sel_exhaust ~> exhaust_sel
|
blanchet@57983
|
1906 |
sel_map ~> map_sel
|
blanchet@57983
|
1907 |
sel_set ~> set_sel
|
blanchet@57983
|
1908 |
sel_split ~> split_sel
|
blanchet@57983
|
1909 |
sel_split_asm ~> split_sel_asm
|
blanchet@57983
|
1910 |
strong_coinduct ~> coinduct_strong
|
blanchet@57983
|
1911 |
weak_case_cong ~> case_cong_weak
|
blanchet@57983
|
1912 |
INCOMPATIBILITY.
|
blanchet@58192
|
1913 |
- The "no_code" option to "free_constructors", "datatype_new", and
|
blanchet@58192
|
1914 |
"codatatype" has been renamed "plugins del: code".
|
blanchet@58192
|
1915 |
INCOMPATIBILITY.
|
blanchet@58044
|
1916 |
- The rules "set_empty" have been removed. They are easy
|
blanchet@58044
|
1917 |
consequences of other set rules "by auto".
|
blanchet@58044
|
1918 |
INCOMPATIBILITY.
|
blanchet@58044
|
1919 |
- The rule "set_cases" is now registered with the "[cases set]"
|
blanchet@57990
|
1920 |
attribute. This can influence the behavior of the "cases" proof
|
blanchet@57990
|
1921 |
method when more than one case rule is applicable (e.g., an
|
blanchet@57990
|
1922 |
assumption is of the form "w : set ws" and the method "cases w"
|
blanchet@57990
|
1923 |
is invoked). The solution is to specify the case rule explicitly
|
blanchet@57990
|
1924 |
(e.g. "cases w rule: widget.exhaust").
|
blanchet@57990
|
1925 |
INCOMPATIBILITY.
|
blanchet@59675
|
1926 |
- Renamed theories:
|
blanchet@59675
|
1927 |
BNF_Comp ~> BNF_Composition
|
blanchet@59675
|
1928 |
BNF_FP_Base ~> BNF_Fixpoint_Base
|
blanchet@59675
|
1929 |
BNF_GFP ~> BNF_Greatest_Fixpoint
|
blanchet@59675
|
1930 |
BNF_LFP ~> BNF_Least_Fixpoint
|
blanchet@59675
|
1931 |
BNF_Constructions_on_Wellorders ~> BNF_Wellorder_Constructions
|
blanchet@59675
|
1932 |
Cardinals/Constructions_on_Wellorders ~> Cardinals/Wellorder_Constructions
|
blanchet@59675
|
1933 |
INCOMPATIBILITY.
|
wenzelm@60115
|
1934 |
- Lifting and Transfer setup for basic HOL types sum and prod (also
|
wenzelm@60115
|
1935 |
option) is now performed by the BNF package. Theories Lifting_Sum,
|
wenzelm@60115
|
1936 |
Lifting_Product and Lifting_Option from Main became obsolete and
|
wenzelm@60115
|
1937 |
were removed. Changed definitions of the relators rel_prod and
|
wenzelm@60115
|
1938 |
rel_sum (using inductive).
|
traytel@60112
|
1939 |
INCOMPATIBILITY: use rel_prod.simps and rel_sum.simps instead
|
wenzelm@60115
|
1940 |
of rel_prod_def and rel_sum_def.
|
wenzelm@60115
|
1941 |
Minor INCOMPATIBILITY: (rarely used by name) transfer theorem names
|
wenzelm@60115
|
1942 |
changed (e.g. map_prod_transfer ~> prod.map_transfer).
|
wenzelm@60261
|
1943 |
- Parametricity theorems for map functions, relators, set functions,
|
wenzelm@60261
|
1944 |
constructors, case combinators, discriminators, selectors and
|
wenzelm@60261
|
1945 |
(co)recursors are automatically proved and registered as transfer
|
wenzelm@60261
|
1946 |
rules.
|
blanchet@57983
|
1947 |
|
blanchet@57983
|
1948 |
* Old datatype package:
|
blanchet@58310
|
1949 |
- The old 'datatype' command has been renamed 'old_datatype', and
|
blanchet@58374
|
1950 |
'rep_datatype' has been renamed 'old_rep_datatype'. They are
|
blanchet@58374
|
1951 |
provided by "~~/src/HOL/Library/Old_Datatype.thy". See
|
blanchet@58310
|
1952 |
'isabelle doc datatypes' for information on porting.
|
blanchet@58374
|
1953 |
INCOMPATIBILITY.
|
blanchet@57983
|
1954 |
- Renamed theorems:
|
blanchet@57983
|
1955 |
weak_case_cong ~> case_cong_weak
|
blanchet@57983
|
1956 |
INCOMPATIBILITY.
|
blanchet@58374
|
1957 |
- Renamed theory:
|
blanchet@58374
|
1958 |
~~/src/HOL/Datatype.thy ~> ~~/src/HOL/Library/Old_Datatype.thy
|
blanchet@58374
|
1959 |
INCOMPATIBILITY.
|
blanchet@57983
|
1960 |
|
blanchet@59039
|
1961 |
* Nitpick:
|
wenzelm@60011
|
1962 |
- Fixed soundness bug related to the strict and non-strict subset
|
blanchet@59039
|
1963 |
operations.
|
blanchet@59039
|
1964 |
|
blanchet@57737
|
1965 |
* Sledgehammer:
|
blanchet@59511
|
1966 |
- CVC4 is now included with Isabelle instead of CVC3 and run by
|
blanchet@59511
|
1967 |
default.
|
blanchet@59965
|
1968 |
- Z3 is now always enabled by default, now that it is fully open
|
blanchet@59965
|
1969 |
source. The "z3_non_commercial" option is discontinued.
|
blanchet@57737
|
1970 |
- Minimization is now always enabled by default.
|
wenzelm@60011
|
1971 |
Removed sub-command:
|
blanchet@57737
|
1972 |
min
|
blanchet@59967
|
1973 |
- Proof reconstruction, both one-liners and Isar, has been
|
blanchet@59039
|
1974 |
dramatically improved.
|
blanchet@59039
|
1975 |
- Improved support for CVC4 and veriT.
|
blanchet@57737
|
1976 |
|
blanchet@58062
|
1977 |
* Old and new SMT modules:
|
blanchet@58067
|
1978 |
- The old 'smt' method has been renamed 'old_smt' and moved to
|
wenzelm@59569
|
1979 |
'src/HOL/Library/Old_SMT.thy'. It is provided for compatibility,
|
wenzelm@59569
|
1980 |
until applications have been ported to use the new 'smt' method. For
|
wenzelm@59569
|
1981 |
the method to work, an older version of Z3 (e.g. Z3 3.2 or 4.0) must
|
wenzelm@59569
|
1982 |
be installed, and the environment variable "OLD_Z3_SOLVER" must
|
wenzelm@59569
|
1983 |
point to it.
|
blanchet@58062
|
1984 |
INCOMPATIBILITY.
|
blanchet@58067
|
1985 |
- The 'smt2' method has been renamed 'smt'.
|
blanchet@58060
|
1986 |
INCOMPATIBILITY.
|
wenzelm@59569
|
1987 |
- New option 'smt_reconstruction_step_timeout' to limit the
|
wenzelm@59569
|
1988 |
reconstruction time of Z3 proof steps in the new 'smt' method.
|
boehmes@59216
|
1989 |
- New option 'smt_statistics' to display statistics of the new 'smt'
|
boehmes@59216
|
1990 |
method, especially runtime statistics of Z3 proof reconstruction.
|
blanchet@58060
|
1991 |
|
wenzelm@60261
|
1992 |
* Lifting: command 'lift_definition' allows to execute lifted constants
|
wenzelm@60261
|
1993 |
that have as a return type a datatype containing a subtype. This
|
wenzelm@60261
|
1994 |
overcomes long-time limitations in the area of code generation and
|
wenzelm@60261
|
1995 |
lifting, and avoids tedious workarounds.
|
kuncar@60258
|
1996 |
|
wenzelm@60010
|
1997 |
* Command and antiquotation "value" provide different evaluation slots
|
wenzelm@60010
|
1998 |
(again), where the previous strategy (NBE after ML) serves as default.
|
wenzelm@60010
|
1999 |
Minor INCOMPATIBILITY.
|
wenzelm@60010
|
2000 |
|
wenzelm@60010
|
2001 |
* Add NO_MATCH-simproc, allows to check for syntactic non-equality.
|
wenzelm@60010
|
2002 |
|
wenzelm@60010
|
2003 |
* field_simps: Use NO_MATCH-simproc for distribution rules, to avoid
|
wenzelm@60010
|
2004 |
non-termination in case of distributing a division. With this change
|
wenzelm@60010
|
2005 |
field_simps is in some cases slightly less powerful, if it fails try to
|
wenzelm@60010
|
2006 |
add algebra_simps, or use divide_simps. Minor INCOMPATIBILITY.
|
wenzelm@60010
|
2007 |
|
wenzelm@60010
|
2008 |
* Separate class no_zero_divisors has been given up in favour of fully
|
wenzelm@60010
|
2009 |
algebraic semiring_no_zero_divisors. INCOMPATIBILITY.
|
wenzelm@60010
|
2010 |
|
wenzelm@60010
|
2011 |
* Class linordered_semidom really requires no zero divisors.
|
wenzelm@60010
|
2012 |
INCOMPATIBILITY.
|
wenzelm@60010
|
2013 |
|
wenzelm@60010
|
2014 |
* Classes division_ring, field and linordered_field always demand
|
wenzelm@60010
|
2015 |
"inverse 0 = 0". Given up separate classes division_ring_inverse_zero,
|
wenzelm@60010
|
2016 |
field_inverse_zero and linordered_field_inverse_zero. INCOMPATIBILITY.
|
wenzelm@60010
|
2017 |
|
wenzelm@60010
|
2018 |
* Classes cancel_ab_semigroup_add / cancel_monoid_add specify explicit
|
wenzelm@60010
|
2019 |
additive inverse operation. INCOMPATIBILITY.
|
wenzelm@60010
|
2020 |
|
lp15@60020
|
2021 |
* Complex powers and square roots. The functions "ln" and "powr" are now
|
wenzelm@60025
|
2022 |
overloaded for types real and complex, and 0 powr y = 0 by definition.
|
wenzelm@60025
|
2023 |
INCOMPATIBILITY: type constraints may be necessary.
|
lp15@60020
|
2024 |
|
wenzelm@60010
|
2025 |
* The functions "sin" and "cos" are now defined for any type of sort
|
wenzelm@60010
|
2026 |
"{real_normed_algebra_1,banach}" type, so in particular on "real" and
|
wenzelm@60010
|
2027 |
"complex" uniformly. Minor INCOMPATIBILITY: type constraints may be
|
wenzelm@60010
|
2028 |
needed.
|
wenzelm@60010
|
2029 |
|
wenzelm@60010
|
2030 |
* New library of properties of the complex transcendental functions sin,
|
wenzelm@60010
|
2031 |
cos, tan, exp, Ln, Arctan, Arcsin, Arccos. Ported from HOL Light.
|
wenzelm@60010
|
2032 |
|
wenzelm@60010
|
2033 |
* The factorial function, "fact", now has type "nat => 'a" (of a sort
|
wenzelm@60010
|
2034 |
that admits numeric types including nat, int, real and complex.
|
wenzelm@60010
|
2035 |
INCOMPATIBILITY: an expression such as "fact 3 = 6" may require a type
|
wenzelm@60010
|
2036 |
constraint, and the combination "real (fact k)" is likely to be
|
wenzelm@60010
|
2037 |
unsatisfactory. If a type conversion is still necessary, then use
|
wenzelm@60010
|
2038 |
"of_nat (fact k)" or "real_of_nat (fact k)".
|
wenzelm@60010
|
2039 |
|
wenzelm@60010
|
2040 |
* Removed functions "natfloor" and "natceiling", use "nat o floor" and
|
wenzelm@60010
|
2041 |
"nat o ceiling" instead. A few of the lemmas have been retained and
|
wenzelm@60010
|
2042 |
adapted: in their names "natfloor"/"natceiling" has been replaced by
|
wenzelm@60010
|
2043 |
"nat_floor"/"nat_ceiling".
|
wenzelm@60010
|
2044 |
|
wenzelm@60010
|
2045 |
* Qualified some duplicated fact names required for boostrapping the
|
wenzelm@60010
|
2046 |
type class hierarchy:
|
wenzelm@60010
|
2047 |
ab_add_uminus_conv_diff ~> diff_conv_add_uminus
|
wenzelm@60010
|
2048 |
field_inverse_zero ~> inverse_zero
|
wenzelm@60010
|
2049 |
field_divide_inverse ~> divide_inverse
|
wenzelm@60010
|
2050 |
field_inverse ~> left_inverse
|
wenzelm@60010
|
2051 |
Minor INCOMPATIBILITY.
|
wenzelm@60010
|
2052 |
|
wenzelm@60010
|
2053 |
* Eliminated fact duplicates:
|
wenzelm@60010
|
2054 |
mult_less_imp_less_right ~> mult_right_less_imp_less
|
wenzelm@60010
|
2055 |
mult_less_imp_less_left ~> mult_left_less_imp_less
|
wenzelm@60010
|
2056 |
Minor INCOMPATIBILITY.
|
wenzelm@60010
|
2057 |
|
wenzelm@60010
|
2058 |
* Fact consolidation: even_less_0_iff is subsumed by
|
wenzelm@60010
|
2059 |
double_add_less_zero_iff_single_add_less_zero (simp by default anyway).
|
wenzelm@60010
|
2060 |
|
wenzelm@60010
|
2061 |
* Generalized and consolidated some theorems concerning divsibility:
|
wenzelm@60010
|
2062 |
dvd_reduce ~> dvd_add_triv_right_iff
|
wenzelm@60010
|
2063 |
dvd_plus_eq_right ~> dvd_add_right_iff
|
wenzelm@60010
|
2064 |
dvd_plus_eq_left ~> dvd_add_left_iff
|
wenzelm@60010
|
2065 |
Minor INCOMPATIBILITY.
|
wenzelm@60010
|
2066 |
|
wenzelm@60010
|
2067 |
* "even" and "odd" are mere abbreviations for "2 dvd _" and "~ 2 dvd _"
|
wenzelm@60010
|
2068 |
and part of theory Main.
|
wenzelm@60010
|
2069 |
even_def ~> even_iff_mod_2_eq_zero
|
wenzelm@60010
|
2070 |
INCOMPATIBILITY.
|
wenzelm@60010
|
2071 |
|
wenzelm@60010
|
2072 |
* Lemma name consolidation: divide_Numeral1 ~> divide_numeral_1. Minor
|
wenzelm@60010
|
2073 |
INCOMPATIBILITY.
|
wenzelm@60010
|
2074 |
|
wenzelm@60010
|
2075 |
* Bootstrap of listsum as special case of abstract product over lists.
|
wenzelm@60010
|
2076 |
Fact rename:
|
wenzelm@60010
|
2077 |
listsum_def ~> listsum.eq_foldr
|
wenzelm@60010
|
2078 |
INCOMPATIBILITY.
|
wenzelm@60010
|
2079 |
|
wenzelm@60010
|
2080 |
* Product over lists via constant "listprod".
|
wenzelm@60010
|
2081 |
|
wenzelm@60010
|
2082 |
* Theory List: renamed drop_Suc_conv_tl and nth_drop' to
|
wenzelm@60010
|
2083 |
Cons_nth_drop_Suc.
|
nipkow@58247
|
2084 |
|
Andreas@58626
|
2085 |
* New infrastructure for compiling, running, evaluating and testing
|
wenzelm@59569
|
2086 |
generated code in target languages in HOL/Library/Code_Test. See
|
wenzelm@59569
|
2087 |
HOL/Codegenerator_Test/Code_Test* for examples.
|
wenzelm@58008
|
2088 |
|
wenzelm@60010
|
2089 |
* Library/Multiset:
|
blanchet@59813
|
2090 |
- Introduced "replicate_mset" operation.
|
blanchet@59813
|
2091 |
- Introduced alternative characterizations of the multiset ordering in
|
blanchet@59813
|
2092 |
"Library/Multiset_Order".
|
blanchet@59958
|
2093 |
- Renamed multiset ordering:
|
blanchet@59958
|
2094 |
<# ~> #<#
|
blanchet@59958
|
2095 |
<=# ~> #<=#
|
blanchet@59958
|
2096 |
\<subset># ~> #\<subset>#
|
blanchet@59958
|
2097 |
\<subseteq># ~> #\<subseteq>#
|
blanchet@59958
|
2098 |
INCOMPATIBILITY.
|
blanchet@59986
|
2099 |
- Introduced abbreviations for ill-named multiset operations:
|
blanchet@59986
|
2100 |
<#, \<subset># abbreviate < (strict subset)
|
blanchet@59986
|
2101 |
<=#, \<le>#, \<subseteq># abbreviate <= (subset or equal)
|
blanchet@59986
|
2102 |
INCOMPATIBILITY.
|
blanchet@59813
|
2103 |
- Renamed
|
blanchet@59813
|
2104 |
in_multiset_of ~> in_multiset_in_set
|
nipkow@59998
|
2105 |
Multiset.fold ~> fold_mset
|
nipkow@59998
|
2106 |
Multiset.filter ~> filter_mset
|
blanchet@59813
|
2107 |
INCOMPATIBILITY.
|
nipkow@59949
|
2108 |
- Removed mcard, is equal to size.
|
blanchet@59813
|
2109 |
- Added attributes:
|
blanchet@59813
|
2110 |
image_mset.id [simp]
|
blanchet@59813
|
2111 |
image_mset_id [simp]
|
blanchet@59813
|
2112 |
elem_multiset_of_set [simp, intro]
|
blanchet@59813
|
2113 |
comp_fun_commute_plus_mset [simp]
|
blanchet@59813
|
2114 |
comp_fun_commute.fold_mset_insert [OF comp_fun_commute_plus_mset, simp]
|
blanchet@59813
|
2115 |
in_mset_fold_plus_iff [iff]
|
blanchet@59813
|
2116 |
set_of_Union_mset [simp]
|
blanchet@59813
|
2117 |
in_Union_mset_iff [iff]
|
blanchet@59813
|
2118 |
INCOMPATIBILITY.
|
blanchet@59813
|
2119 |
|
wenzelm@60010
|
2120 |
* Library/Sum_of_Squares: simplified and improved "sos" method. Always
|
wenzelm@60010
|
2121 |
use local CSDP executable, which is much faster than the NEOS server.
|
wenzelm@60010
|
2122 |
The "sos_cert" functionality is invoked as "sos" with additional
|
wenzelm@60010
|
2123 |
argument. Minor INCOMPATIBILITY.
|
wenzelm@60010
|
2124 |
|
wenzelm@60010
|
2125 |
* HOL-Decision_Procs: New counterexample generator quickcheck
|
wenzelm@60010
|
2126 |
[approximation] for inequalities of transcendental functions. Uses
|
wenzelm@60010
|
2127 |
hardware floating point arithmetic to randomly discover potential
|
wenzelm@60011
|
2128 |
counterexamples. Counterexamples are certified with the "approximation"
|
wenzelm@60010
|
2129 |
method. See HOL/Decision_Procs/ex/Approximation_Quickcheck_Ex.thy for
|
wenzelm@60010
|
2130 |
examples.
|
immler@58990
|
2131 |
|
hoelzl@59354
|
2132 |
* HOL-Probability: Reworked measurability prover
|
wenzelm@60011
|
2133 |
- applies destructor rules repeatedly
|
hoelzl@59354
|
2134 |
- removed application splitting (replaced by destructor rule)
|
wenzelm@59569
|
2135 |
- added congruence rules to rewrite measure spaces under the sets
|
wenzelm@59569
|
2136 |
projection
|
wenzelm@59569
|
2137 |
|
wenzelm@60010
|
2138 |
* New proof method "rewrite" (in theory ~~/src/HOL/Library/Rewrite) for
|
wenzelm@60010
|
2139 |
single-step rewriting with subterm selection based on patterns.
|
wenzelm@60010
|
2140 |
|
wenzelm@58630
|
2141 |
|
blanchet@58066
|
2142 |
*** ML ***
|
blanchet@58066
|
2143 |
|
wenzelm@60010
|
2144 |
* Subtle change of name space policy: undeclared entries are now
|
wenzelm@60010
|
2145 |
considered inaccessible, instead of accessible via the fully-qualified
|
wenzelm@60010
|
2146 |
internal name. This mainly affects Name_Space.intern (and derivatives),
|
wenzelm@60010
|
2147 |
which may produce an unexpected Long_Name.hidden prefix. Note that
|
wenzelm@60011
|
2148 |
contemporary applications use the strict Name_Space.check (and
|
wenzelm@60010
|
2149 |
derivatives) instead, which is not affected by the change. Potential
|
wenzelm@60010
|
2150 |
INCOMPATIBILITY in rare applications of Name_Space.intern.
|
wenzelm@59951
|
2151 |
|
wenzelm@60096
|
2152 |
* Subtle change of error semantics of Toplevel.proof_of: regular user
|
wenzelm@60096
|
2153 |
ERROR instead of internal Toplevel.UNDEF.
|
wenzelm@60096
|
2154 |
|
wenzelm@59951
|
2155 |
* Basic combinators map, fold, fold_map, split_list, apply are available
|
wenzelm@59951
|
2156 |
as parameterized antiquotations, e.g. @{map 4} for lists of quadruples.
|
wenzelm@59951
|
2157 |
|
wenzelm@59951
|
2158 |
* Renamed "pairself" to "apply2", in accordance to @{apply 2}.
|
wenzelm@59951
|
2159 |
INCOMPATIBILITY.
|
wenzelm@59951
|
2160 |
|
wenzelm@59951
|
2161 |
* Former combinators NAMED_CRITICAL and CRITICAL for central critical
|
wenzelm@59951
|
2162 |
sections have been discontinued, in favour of the more elementary
|
wenzelm@59951
|
2163 |
Multithreading.synchronized and its high-level derivative
|
wenzelm@59951
|
2164 |
Synchronized.var (which is usually sufficient in applications). Subtle
|
wenzelm@59951
|
2165 |
INCOMPATIBILITY: synchronized access needs to be atomic and cannot be
|
wenzelm@59951
|
2166 |
nested.
|
wenzelm@59951
|
2167 |
|
wenzelm@60010
|
2168 |
* Synchronized.value (ML) is actually synchronized (as in Scala): subtle
|
wenzelm@60010
|
2169 |
change of semantics with minimal potential for INCOMPATIBILITY.
|
wenzelm@59899
|
2170 |
|
wenzelm@59621
|
2171 |
* The main operations to certify logical entities are Thm.ctyp_of and
|
wenzelm@59621
|
2172 |
Thm.cterm_of with a local context; old-style global theory variants are
|
wenzelm@59621
|
2173 |
available as Thm.global_ctyp_of and Thm.global_cterm_of.
|
wenzelm@59621
|
2174 |
INCOMPATIBILITY.
|
wenzelm@59621
|
2175 |
|
wenzelm@59582
|
2176 |
* Elementary operations in module Thm are no longer pervasive.
|
wenzelm@59582
|
2177 |
INCOMPATIBILITY, need to use qualified Thm.prop_of, Thm.cterm_of,
|
wenzelm@59582
|
2178 |
Thm.term_of etc.
|
wenzelm@59582
|
2179 |
|
wenzelm@58963
|
2180 |
* Proper context for various elementary tactics: assume_tac,
|
wenzelm@59498
|
2181 |
resolve_tac, eresolve_tac, dresolve_tac, forward_tac, match_tac,
|
wenzelm@59498
|
2182 |
compose_tac, Splitter.split_tac etc. INCOMPATIBILITY.
|
wenzelm@58956
|
2183 |
|
blanchet@58066
|
2184 |
* Tactical PARALLEL_ALLGOALS is the most common way to refer to
|
blanchet@58066
|
2185 |
PARALLEL_GOALS.
|
blanchet@58066
|
2186 |
|
wenzelm@59564
|
2187 |
* Goal.prove_multi is superseded by the fully general Goal.prove_common,
|
wenzelm@59564
|
2188 |
which also allows to specify a fork priority.
|
wenzelm@59564
|
2189 |
|
wenzelm@59936
|
2190 |
* Antiquotation @{command_spec "COMMAND"} is superseded by
|
wenzelm@59936
|
2191 |
@{command_keyword COMMAND} (usually without quotes and with PIDE
|
wenzelm@59936
|
2192 |
markup). Minor INCOMPATIBILITY.
|
wenzelm@59936
|
2193 |
|
wenzelm@60010
|
2194 |
* Cartouches within ML sources are turned into values of type
|
wenzelm@60010
|
2195 |
Input.source (with formal position information).
|
wenzelm@60010
|
2196 |
|
blanchet@58066
|
2197 |
|
wenzelm@58610
|
2198 |
*** System ***
|
wenzelm@58610
|
2199 |
|
wenzelm@59951
|
2200 |
* The Isabelle tool "update_cartouches" changes theory files to use
|
wenzelm@59951
|
2201 |
cartouches instead of old-style {* verbatim *} or `alt_string` tokens.
|
wenzelm@59951
|
2202 |
|
wenzelm@60108
|
2203 |
* The Isabelle tool "build" provides new options -X, -k, -x.
|
wenzelm@59951
|
2204 |
|
wenzelm@59951
|
2205 |
* Discontinued old-fashioned "codegen" tool. Code generation can always
|
wenzelm@59951
|
2206 |
be externally triggered using an appropriate ROOT file plus a
|
wenzelm@59951
|
2207 |
corresponding theory. Parametrization is possible using environment
|
wenzelm@59951
|
2208 |
variables, or ML snippets in the most extreme cases. Minor
|
wenzelm@59951
|
2209 |
INCOMPATIBILITY.
|
wenzelm@58842
|
2210 |
|
wenzelm@59200
|
2211 |
* JVM system property "isabelle.threads" determines size of Scala thread
|
wenzelm@59200
|
2212 |
pool, like Isabelle system option "threads" for ML.
|
wenzelm@59200
|
2213 |
|
wenzelm@59201
|
2214 |
* JVM system property "isabelle.laf" determines the default Swing
|
wenzelm@59201
|
2215 |
look-and-feel, via internal class name or symbolic name as in the jEdit
|
wenzelm@59201
|
2216 |
menu Global Options / Appearance.
|
wenzelm@59201
|
2217 |
|
wenzelm@59951
|
2218 |
* Support for Proof General and Isar TTY loop has been discontinued.
|
wenzelm@60010
|
2219 |
Minor INCOMPATIBILITY, use standard PIDE infrastructure instead.
|
wenzelm@59891
|
2220 |
|
wenzelm@58610
|
2221 |
|
wenzelm@57695
|
2222 |
|
wenzelm@57452
|
2223 |
New in Isabelle2014 (August 2014)
|
wenzelm@57452
|
2224 |
---------------------------------
|
wenzelm@54055
|
2225 |
|
wenzelm@54702
|
2226 |
*** General ***
|
wenzelm@54702
|
2227 |
|
wenzelm@57452
|
2228 |
* Support for official Standard ML within the Isabelle context.
|
wenzelm@57452
|
2229 |
Command 'SML_file' reads and evaluates the given Standard ML file.
|
wenzelm@57452
|
2230 |
Toplevel bindings are stored within the theory context; the initial
|
wenzelm@57452
|
2231 |
environment is restricted to the Standard ML implementation of
|
wenzelm@57452
|
2232 |
Poly/ML, without the add-ons of Isabelle/ML. Commands 'SML_import'
|
wenzelm@57452
|
2233 |
and 'SML_export' allow to exchange toplevel bindings between the two
|
wenzelm@57452
|
2234 |
separate environments. See also ~~/src/Tools/SML/Examples.thy for
|
wenzelm@57452
|
2235 |
some examples.
|
wenzelm@56499
|
2236 |
|
wenzelm@57504
|
2237 |
* Standard tactics and proof methods such as "clarsimp", "auto" and
|
wenzelm@57504
|
2238 |
"safe" now preserve equality hypotheses "x = expr" where x is a free
|
wenzelm@57504
|
2239 |
variable. Locale assumptions and chained facts containing "x"
|
wenzelm@57504
|
2240 |
continue to be useful. The new method "hypsubst_thin" and the
|
wenzelm@57504
|
2241 |
configuration option "hypsubst_thin" (within the attribute name space)
|
wenzelm@57504
|
2242 |
restore the previous behavior. INCOMPATIBILITY, especially where
|
wenzelm@57504
|
2243 |
induction is done after these methods or when the names of free and
|
wenzelm@57504
|
2244 |
bound variables clash. As first approximation, old proofs may be
|
wenzelm@57504
|
2245 |
repaired by "using [[hypsubst_thin = true]]" in the critical spot.
|
wenzelm@57504
|
2246 |
|
wenzelm@56232
|
2247 |
* More static checking of proof methods, which allows the system to
|
wenzelm@56232
|
2248 |
form a closure over the concrete syntax. Method arguments should be
|
wenzelm@56232
|
2249 |
processed in the original proof context as far as possible, before
|
wenzelm@56232
|
2250 |
operating on the goal state. In any case, the standard discipline for
|
wenzelm@56232
|
2251 |
subgoal-addressing needs to be observed: no subgoals or a subgoal
|
wenzelm@56232
|
2252 |
number that is out of range produces an empty result sequence, not an
|
wenzelm@56232
|
2253 |
exception. Potential INCOMPATIBILITY for non-conformant tactical
|
wenzelm@56232
|
2254 |
proof tools.
|
wenzelm@56232
|
2255 |
|
wenzelm@57452
|
2256 |
* Lexical syntax (inner and outer) supports text cartouches with
|
wenzelm@57452
|
2257 |
arbitrary nesting, and without escapes of quotes etc. The Prover IDE
|
wenzelm@57452
|
2258 |
supports input via ` (backquote).
|
wenzelm@57452
|
2259 |
|
wenzelm@57452
|
2260 |
* The outer syntax categories "text" (for formal comments and document
|
wenzelm@57452
|
2261 |
markup commands) and "altstring" (for literal fact references) allow
|
wenzelm@57452
|
2262 |
cartouches as well, in addition to the traditional mix of quotations.
|
wenzelm@57452
|
2263 |
|
wenzelm@57452
|
2264 |
* Syntax of document antiquotation @{rail} now uses \<newline> instead
|
wenzelm@57452
|
2265 |
of "\\", to avoid the optical illusion of escaped backslash within
|
wenzelm@57491
|
2266 |
string token. General renovation of its syntax using text cartouches.
|
wenzelm@57452
|
2267 |
Minor INCOMPATIBILITY.
|
wenzelm@57452
|
2268 |
|
wenzelm@57452
|
2269 |
* Discontinued legacy_isub_isup, which was a temporary workaround for
|
wenzelm@57452
|
2270 |
Isabelle/ML in Isabelle2013-1. The prover process no longer accepts
|
wenzelm@57452
|
2271 |
old identifier syntax with \<^isub> or \<^isup>. Potential
|
wenzelm@57452
|
2272 |
INCOMPATIBILITY.
|
wenzelm@57452
|
2273 |
|
wenzelm@57452
|
2274 |
* Document antiquotation @{url} produces markup for the given URL,
|
wenzelm@57452
|
2275 |
which results in an active hyperlink within the text.
|
wenzelm@57452
|
2276 |
|
wenzelm@57452
|
2277 |
* Document antiquotation @{file_unchecked} is like @{file}, but does
|
wenzelm@57452
|
2278 |
not check existence within the file-system.
|
wenzelm@57452
|
2279 |
|
wenzelm@57452
|
2280 |
* Updated and extended manuals: codegen, datatypes, implementation,
|
wenzelm@57452
|
2281 |
isar-ref, jedit, system.
|
wenzelm@57423
|
2282 |
|
wenzelm@54702
|
2283 |
|
wenzelm@54533
|
2284 |
*** Prover IDE -- Isabelle/Scala/jEdit ***
|
wenzelm@54533
|
2285 |
|
wenzelm@57650
|
2286 |
* Improved Document panel: simplified interaction where every single
|
wenzelm@57452
|
2287 |
mouse click (re)opens document via desktop environment or as jEdit
|
wenzelm@57452
|
2288 |
buffer.
|
wenzelm@57452
|
2289 |
|
wenzelm@57452
|
2290 |
* Support for Navigator plugin (with toolbar buttons), with connection
|
wenzelm@57452
|
2291 |
to PIDE hyperlinks.
|
wenzelm@57452
|
2292 |
|
wenzelm@57452
|
2293 |
* Auxiliary files ('ML_file' etc.) are managed by the Prover IDE.
|
wenzelm@57452
|
2294 |
Open text buffers take precedence over copies within the file-system.
|
wenzelm@57452
|
2295 |
|
wenzelm@57452
|
2296 |
* Improved support for Isabelle/ML, with jEdit mode "isabelle-ml" for
|
wenzelm@57452
|
2297 |
auxiliary ML files.
|
wenzelm@57423
|
2298 |
|
wenzelm@57423
|
2299 |
* Improved syntactic and semantic completion mechanism, with simple
|
wenzelm@57423
|
2300 |
templates, completion language context, name-space completion,
|
wenzelm@57423
|
2301 |
file-name completion, spell-checker completion.
|
wenzelm@57423
|
2302 |
|
wenzelm@57423
|
2303 |
* Refined GUI popup for completion: more robust key/mouse event
|
wenzelm@57423
|
2304 |
handling and propagation to enclosing text area -- avoid loosing
|
wenzelm@57423
|
2305 |
keystrokes with slow / remote graphics displays.
|
wenzelm@57423
|
2306 |
|
wenzelm@57833
|
2307 |
* Completion popup supports both ENTER and TAB (default) to select an
|
wenzelm@57833
|
2308 |
item, depending on Isabelle options.
|
wenzelm@57833
|
2309 |
|
wenzelm@57423
|
2310 |
* Refined insertion of completion items wrt. jEdit text: multiple
|
wenzelm@57423
|
2311 |
selections, rectangular selections, rectangular selection as "tall
|
wenzelm@57423
|
2312 |
caret".
|
wenzelm@56342
|
2313 |
|
wenzelm@56580
|
2314 |
* Integrated spell-checker for document text, comments etc. with
|
wenzelm@57423
|
2315 |
completion popup and context-menu.
|
wenzelm@56554
|
2316 |
|
wenzelm@56879
|
2317 |
* More general "Query" panel supersedes "Find" panel, with GUI access
|
wenzelm@56879
|
2318 |
to commands 'find_theorems' and 'find_consts', as well as print
|
wenzelm@56879
|
2319 |
operations for the context. Minor incompatibility in keyboard
|
wenzelm@56879
|
2320 |
shortcuts etc.: replace action isabelle-find by isabelle-query.
|
wenzelm@56761
|
2321 |
|
wenzelm@56901
|
2322 |
* Search field for all output panels ("Output", "Query", "Info" etc.)
|
wenzelm@56901
|
2323 |
to highlight text via regular expression.
|
wenzelm@56901
|
2324 |
|
wenzelm@54881
|
2325 |
* Option "jedit_print_mode" (see also "Plugin Options / Isabelle /
|
wenzelm@54881
|
2326 |
General") allows to specify additional print modes for the prover
|
wenzelm@54881
|
2327 |
process, without requiring old-fashioned command-line invocation of
|
wenzelm@54881
|
2328 |
"isabelle jedit -m MODE".
|
wenzelm@54881
|
2329 |
|
wenzelm@56505
|
2330 |
* More support for remote files (e.g. http) using standard Java
|
wenzelm@56505
|
2331 |
networking operations instead of jEdit virtual file-systems.
|
wenzelm@56505
|
2332 |
|
wenzelm@57822
|
2333 |
* Empty editors buffers that are no longer required (e.g.\ via theory
|
wenzelm@57822
|
2334 |
imports) are automatically removed from the document model.
|
wenzelm@57822
|
2335 |
|
wenzelm@57869
|
2336 |
* Improved monitor panel.
|
wenzelm@57869
|
2337 |
|
wenzelm@56838
|
2338 |
* Improved Console/Scala plugin: more uniform scala.Console output,
|
wenzelm@56838
|
2339 |
more robust treatment of threads and interrupts.
|
wenzelm@56838
|
2340 |
|
wenzelm@56939
|
2341 |
* Improved management of dockable windows: clarified keyboard focus
|
wenzelm@56939
|
2342 |
and window placement wrt. main editor view; optional menu item to
|
wenzelm@56939
|
2343 |
"Detach" a copy where this makes sense.
|
wenzelm@56939
|
2344 |
|
wenzelm@57452
|
2345 |
* New Simplifier Trace panel provides an interactive view of the
|
wenzelm@57594
|
2346 |
simplification process, enabled by the "simp_trace_new" attribute
|
wenzelm@57452
|
2347 |
within the context.
|
wenzelm@57452
|
2348 |
|
wenzelm@57452
|
2349 |
|
wenzelm@55001
|
2350 |
*** Pure ***
|
wenzelm@55001
|
2351 |
|
wenzelm@57504
|
2352 |
* Low-level type-class commands 'classes', 'classrel', 'arities' have
|
wenzelm@57504
|
2353 |
been discontinued to avoid the danger of non-trivial axiomatization
|
wenzelm@57504
|
2354 |
that is not immediately visible. INCOMPATIBILITY, use regular
|
wenzelm@57504
|
2355 |
'instance' command with proof. The required OFCLASS(...) theorem
|
wenzelm@57504
|
2356 |
might be postulated via 'axiomatization' beforehand, or the proof
|
wenzelm@57504
|
2357 |
finished trivially if the underlying class definition is made vacuous
|
wenzelm@57504
|
2358 |
(without any assumptions). See also Isabelle/ML operations
|
wenzelm@57504
|
2359 |
Axclass.class_axiomatization, Axclass.classrel_axiomatization,
|
wenzelm@57504
|
2360 |
Axclass.arity_axiomatization.
|
wenzelm@57504
|
2361 |
|
wenzelm@56245
|
2362 |
* Basic constants of Pure use more conventional names and are always
|
wenzelm@56245
|
2363 |
qualified. Rare INCOMPATIBILITY, but with potentially serious
|
wenzelm@56245
|
2364 |
consequences, notably for tools in Isabelle/ML. The following
|
wenzelm@56245
|
2365 |
renaming needs to be applied:
|
wenzelm@56245
|
2366 |
|
wenzelm@56245
|
2367 |
== ~> Pure.eq
|
wenzelm@56245
|
2368 |
==> ~> Pure.imp
|
wenzelm@56245
|
2369 |
all ~> Pure.all
|
wenzelm@56245
|
2370 |
TYPE ~> Pure.type
|
wenzelm@56245
|
2371 |
dummy_pattern ~> Pure.dummy_pattern
|
wenzelm@56245
|
2372 |
|
wenzelm@56245
|
2373 |
Systematic porting works by using the following theory setup on a
|
wenzelm@56245
|
2374 |
*previous* Isabelle version to introduce the new name accesses for the
|
wenzelm@56245
|
2375 |
old constants:
|
wenzelm@56245
|
2376 |
|
wenzelm@56245
|
2377 |
setup {*
|
wenzelm@56245
|
2378 |
fn thy => thy
|
wenzelm@56245
|
2379 |
|> Sign.root_path
|
wenzelm@56245
|
2380 |
|> Sign.const_alias (Binding.qualify true "Pure" @{binding eq}) "=="
|
wenzelm@56245
|
2381 |
|> Sign.const_alias (Binding.qualify true "Pure" @{binding imp}) "==>"
|
wenzelm@56245
|
2382 |
|> Sign.const_alias (Binding.qualify true "Pure" @{binding all}) "all"
|
wenzelm@56245
|
2383 |
|> Sign.restore_naming thy
|
wenzelm@56245
|
2384 |
*}
|
wenzelm@56245
|
2385 |
|
wenzelm@56245
|
2386 |
Thus ML antiquotations like @{const_name Pure.eq} may be used already.
|
wenzelm@56245
|
2387 |
Later the application is moved to the current Isabelle version, and
|
wenzelm@56245
|
2388 |
the auxiliary aliases are deleted.
|
wenzelm@56245
|
2389 |
|
wenzelm@55143
|
2390 |
* Attributes "where" and "of" allow an optional context of local
|
wenzelm@55143
|
2391 |
variables ('for' declaration): these variables become schematic in the
|
wenzelm@55143
|
2392 |
instantiated theorem.
|
wenzelm@55143
|
2393 |
|
wenzelm@55152
|
2394 |
* Obsolete attribute "standard" has been discontinued (legacy since
|
wenzelm@55152
|
2395 |
Isabelle2012). Potential INCOMPATIBILITY, use explicit 'for' context
|
wenzelm@55152
|
2396 |
where instantiations with schematic variables are intended (for
|
wenzelm@55152
|
2397 |
declaration commands like 'lemmas' or attributes like "of"). The
|
wenzelm@55152
|
2398 |
following temporary definition may help to port old applications:
|
wenzelm@55152
|
2399 |
|
wenzelm@55152
|
2400 |
attribute_setup standard =
|
wenzelm@55152
|
2401 |
"Scan.succeed (Thm.rule_attribute (K Drule.export_without_context))"
|
wenzelm@55152
|
2402 |
|
wenzelm@55001
|
2403 |
* More thorough check of proof context for goal statements and
|
wenzelm@55006
|
2404 |
attributed fact expressions (concerning background theory, declared
|
wenzelm@55006
|
2405 |
hyps). Potential INCOMPATIBILITY, tools need to observe standard
|
wenzelm@55006
|
2406 |
context discipline. See also Assumption.add_assumes and the more
|
wenzelm@55006
|
2407 |
primitive Thm.assume_hyps.
|
wenzelm@55001
|
2408 |
|
wenzelm@55108
|
2409 |
* Inner syntax token language allows regular quoted strings "..."
|
wenzelm@55108
|
2410 |
(only makes sense in practice, if outer syntax is delimited
|
wenzelm@57452
|
2411 |
differently, e.g. via cartouches).
|
wenzelm@57452
|
2412 |
|
wenzelm@57504
|
2413 |
* Command 'print_term_bindings' supersedes 'print_binds' for clarity,
|
wenzelm@57504
|
2414 |
but the latter is retained some time as Proof General legacy.
|
wenzelm@57504
|
2415 |
|
wenzelm@57452
|
2416 |
* Code generator preprocessor: explicit control of simp tracing on a
|
wenzelm@57452
|
2417 |
per-constant basis. See attribute "code_preproc".
|
haftmann@57430
|
2418 |
|
wenzelm@55001
|
2419 |
|
haftmann@54227
|
2420 |
*** HOL ***
|
haftmann@54227
|
2421 |
|
wenzelm@57504
|
2422 |
* Code generator: enforce case of identifiers only for strict target
|
wenzelm@57504
|
2423 |
language requirements. INCOMPATIBILITY.
|
wenzelm@57504
|
2424 |
|
wenzelm@57504
|
2425 |
* Code generator: explicit proof contexts in many ML interfaces.
|
wenzelm@57504
|
2426 |
INCOMPATIBILITY.
|
wenzelm@57504
|
2427 |
|
wenzelm@57504
|
2428 |
* Code generator: minimize exported identifiers by default. Minor
|
wenzelm@57504
|
2429 |
INCOMPATIBILITY.
|
wenzelm@57504
|
2430 |
|
wenzelm@57504
|
2431 |
* Code generation for SML and OCaml: dropped arcane "no_signatures"
|
wenzelm@57504
|
2432 |
option. Minor INCOMPATIBILITY.
|
wenzelm@57504
|
2433 |
|
wenzelm@57504
|
2434 |
* "declare [[code abort: ...]]" replaces "code_abort ...".
|
wenzelm@57504
|
2435 |
INCOMPATIBILITY.
|
wenzelm@57504
|
2436 |
|
wenzelm@57504
|
2437 |
* "declare [[code drop: ...]]" drops all code equations associated
|
wenzelm@57504
|
2438 |
with the given constants.
|
wenzelm@57504
|
2439 |
|
wenzelm@57504
|
2440 |
* Code generations are provided for make, fields, extend and truncate
|
wenzelm@57504
|
2441 |
operations on records.
|
haftmann@57437
|
2442 |
|
wenzelm@57452
|
2443 |
* Command and antiquotation "value" are now hardcoded against nbe and
|
wenzelm@57452
|
2444 |
ML. Minor INCOMPATIBILITY.
|
wenzelm@57452
|
2445 |
|
wenzelm@57504
|
2446 |
* Renamed command 'enriched_type' to 'functor'. INCOMPATIBILITY.
|
wenzelm@57504
|
2447 |
|
wenzelm@57504
|
2448 |
* The symbol "\<newline>" may be used within char or string literals
|
wenzelm@57504
|
2449 |
to represent (Char Nibble0 NibbleA), i.e. ASCII newline.
|
wenzelm@57504
|
2450 |
|
wenzelm@57504
|
2451 |
* Qualified String.implode and String.explode. INCOMPATIBILITY.
|
haftmann@56923
|
2452 |
|
wenzelm@57452
|
2453 |
* Simplifier: Enhanced solver of preconditions of rewrite rules can
|
wenzelm@57452
|
2454 |
now deal with conjunctions. For help with converting proofs, the old
|
wenzelm@57452
|
2455 |
behaviour of the simplifier can be restored like this: declare/using
|
wenzelm@57452
|
2456 |
[[simp_legacy_precond]]. This configuration option will disappear
|
wenzelm@57452
|
2457 |
again in the future. INCOMPATIBILITY.
|
nipkow@56073
|
2458 |
|
wenzelm@55139
|
2459 |
* Simproc "finite_Collect" is no longer enabled by default, due to
|
wenzelm@55139
|
2460 |
spurious crashes and other surprises. Potential INCOMPATIBILITY.
|
wenzelm@55139
|
2461 |
|
wenzelm@57452
|
2462 |
* Moved new (co)datatype package and its dependencies from session
|
wenzelm@57452
|
2463 |
"HOL-BNF" to "HOL". The commands 'bnf', 'wrap_free_constructors',
|
wenzelm@57452
|
2464 |
'datatype_new', 'codatatype', 'primcorec', 'primcorecursive' are now
|
wenzelm@57452
|
2465 |
part of theory "Main".
|
wenzelm@57452
|
2466 |
|
blanchet@55098
|
2467 |
Theory renamings:
|
blanchet@55098
|
2468 |
FunDef.thy ~> Fun_Def.thy (and Fun_Def_Base.thy)
|
blanchet@55098
|
2469 |
Library/Wfrec.thy ~> Wfrec.thy
|
blanchet@55098
|
2470 |
Library/Zorn.thy ~> Zorn.thy
|
blanchet@55098
|
2471 |
Cardinals/Order_Relation.thy ~> Order_Relation.thy
|
blanchet@55098
|
2472 |
Library/Order_Union.thy ~> Cardinals/Order_Union.thy
|
blanchet@55098
|
2473 |
Cardinals/Cardinal_Arithmetic_Base.thy ~> BNF_Cardinal_Arithmetic.thy
|
blanchet@55098
|
2474 |
Cardinals/Cardinal_Order_Relation_Base.thy ~> BNF_Cardinal_Order_Relation.thy
|
blanchet@55098
|
2475 |
Cardinals/Constructions_on_Wellorders_Base.thy ~> BNF_Constructions_on_Wellorders.thy
|
blanchet@55098
|
2476 |
Cardinals/Wellorder_Embedding_Base.thy ~> BNF_Wellorder_Embedding.thy
|
blanchet@55098
|
2477 |
Cardinals/Wellorder_Relation_Base.thy ~> BNF_Wellorder_Relation.thy
|
blanchet@55098
|
2478 |
BNF/Ctr_Sugar.thy ~> Ctr_Sugar.thy
|
blanchet@55098
|
2479 |
BNF/Basic_BNFs.thy ~> Basic_BNFs.thy
|
blanchet@55098
|
2480 |
BNF/BNF_Comp.thy ~> BNF_Comp.thy
|
blanchet@55098
|
2481 |
BNF/BNF_Def.thy ~> BNF_Def.thy
|
blanchet@55098
|
2482 |
BNF/BNF_FP_Base.thy ~> BNF_FP_Base.thy
|
blanchet@55098
|
2483 |
BNF/BNF_GFP.thy ~> BNF_GFP.thy
|
blanchet@55098
|
2484 |
BNF/BNF_LFP.thy ~> BNF_LFP.thy
|
blanchet@55098
|
2485 |
BNF/BNF_Util.thy ~> BNF_Util.thy
|
blanchet@55098
|
2486 |
BNF/Coinduction.thy ~> Coinduction.thy
|
blanchet@55098
|
2487 |
BNF/More_BNFs.thy ~> Library/More_BNFs.thy
|
blanchet@55098
|
2488 |
BNF/Countable_Type.thy ~> Library/Countable_Set_Type.thy
|
blanchet@55098
|
2489 |
BNF/Examples/* ~> BNF_Examples/*
|
wenzelm@57452
|
2490 |
|
blanchet@55098
|
2491 |
New theories:
|
blanchet@55098
|
2492 |
Wellorder_Extension.thy (split from Zorn.thy)
|
blanchet@55098
|
2493 |
Library/Cardinal_Notations.thy
|
traytel@56942
|
2494 |
Library/BNF_Axomatization.thy
|
blanchet@55098
|
2495 |
BNF_Examples/Misc_Primcorec.thy
|
blanchet@55098
|
2496 |
BNF_Examples/Stream_Processor.thy
|
wenzelm@57452
|
2497 |
|
blanchet@55519
|
2498 |
Discontinued theories:
|
blanchet@55098
|
2499 |
BNF/BNF.thy
|
blanchet@55098
|
2500 |
BNF/Equiv_Relations_More.thy
|
wenzelm@57452
|
2501 |
|
wenzelm@57452
|
2502 |
INCOMPATIBILITY.
|
blanchet@55098
|
2503 |
|
blanchet@56118
|
2504 |
* New (co)datatype package:
|
wenzelm@57452
|
2505 |
- Command 'primcorec' is fully implemented.
|
wenzelm@57452
|
2506 |
- Command 'datatype_new' generates size functions ("size_xxx" and
|
wenzelm@57452
|
2507 |
"size") as required by 'fun'.
|
wenzelm@57452
|
2508 |
- BNFs are integrated with the Lifting tool and new-style
|
wenzelm@57452
|
2509 |
(co)datatypes with Transfer.
|
wenzelm@57452
|
2510 |
- Renamed commands:
|
blanchet@55875
|
2511 |
datatype_new_compat ~> datatype_compat
|
blanchet@55875
|
2512 |
primrec_new ~> primrec
|
blanchet@55875
|
2513 |
wrap_free_constructors ~> free_constructors
|
blanchet@55875
|
2514 |
INCOMPATIBILITY.
|
wenzelm@57452
|
2515 |
- The generated constants "xxx_case" and "xxx_rec" have been renamed
|
blanchet@55875
|
2516 |
"case_xxx" and "rec_xxx" (e.g., "prod_case" ~> "case_prod").
|
blanchet@55875
|
2517 |
INCOMPATIBILITY.
|
wenzelm@57452
|
2518 |
- The constant "xxx_(un)fold" and related theorems are no longer
|
wenzelm@57452
|
2519 |
generated. Use "xxx_(co)rec" or define "xxx_(un)fold" manually
|
wenzelm@57452
|
2520 |
using "prim(co)rec".
|
blanchet@55875
|
2521 |
INCOMPATIBILITY.
|
wenzelm@57452
|
2522 |
- No discriminators are generated for nullary constructors by
|
wenzelm@57452
|
2523 |
default, eliminating the need for the odd "=:" syntax.
|
blanchet@57091
|
2524 |
INCOMPATIBILITY.
|
wenzelm@57452
|
2525 |
- No discriminators or selectors are generated by default by
|
blanchet@57094
|
2526 |
"datatype_new", unless custom names are specified or the new
|
blanchet@57094
|
2527 |
"discs_sels" option is passed.
|
blanchet@57094
|
2528 |
INCOMPATIBILITY.
|
blanchet@55875
|
2529 |
|
blanchet@55643
|
2530 |
* Old datatype package:
|
wenzelm@57452
|
2531 |
- The generated theorems "xxx.cases" and "xxx.recs" have been
|
wenzelm@57452
|
2532 |
renamed "xxx.case" and "xxx.rec" (e.g., "sum.cases" ->
|
wenzelm@57452
|
2533 |
"sum.case"). INCOMPATIBILITY.
|
wenzelm@57452
|
2534 |
- The generated constants "xxx_case", "xxx_rec", and "xxx_size" have
|
wenzelm@57452
|
2535 |
been renamed "case_xxx", "rec_xxx", and "size_xxx" (e.g.,
|
wenzelm@57452
|
2536 |
"prod_case" ~> "case_prod"). INCOMPATIBILITY.
|
wenzelm@57452
|
2537 |
|
wenzelm@57452
|
2538 |
* The types "'a list" and "'a option", their set and map functions,
|
wenzelm@57452
|
2539 |
their relators, and their selectors are now produced using the new
|
wenzelm@57452
|
2540 |
BNF-based datatype package.
|
wenzelm@57452
|
2541 |
|
blanchet@55519
|
2542 |
Renamed constants:
|
blanchet@55519
|
2543 |
Option.set ~> set_option
|
blanchet@55519
|
2544 |
Option.map ~> map_option
|
blanchet@55525
|
2545 |
option_rel ~> rel_option
|
wenzelm@57452
|
2546 |
|
blanchet@55519
|
2547 |
Renamed theorems:
|
blanchet@55585
|
2548 |
set_def ~> set_rec[abs_def]
|
blanchet@55519
|
2549 |
map_def ~> map_rec[abs_def]
|
blanchet@55519
|
2550 |
Option.map_def ~> map_option_case[abs_def] (with "case_option" instead of "rec_option")
|
blanchet@56652
|
2551 |
option.recs ~> option.rec
|
blanchet@55524
|
2552 |
list_all2_def ~> list_all2_iff
|
blanchet@55585
|
2553 |
set.simps ~> set_simps (or the slightly different "list.set")
|
blanchet@55519
|
2554 |
map.simps ~> list.map
|
blanchet@55519
|
2555 |
hd.simps ~> list.sel(1)
|
blanchet@55519
|
2556 |
tl.simps ~> list.sel(2-3)
|
blanchet@55519
|
2557 |
the.simps ~> option.sel
|
wenzelm@57452
|
2558 |
|
wenzelm@57452
|
2559 |
INCOMPATIBILITY.
|
blanchet@55519
|
2560 |
|
blanchet@55933
|
2561 |
* The following map functions and relators have been renamed:
|
blanchet@55939
|
2562 |
sum_map ~> map_sum
|
blanchet@55939
|
2563 |
map_pair ~> map_prod
|
blanchet@55944
|
2564 |
prod_rel ~> rel_prod
|
blanchet@55943
|
2565 |
sum_rel ~> rel_sum
|
blanchet@55945
|
2566 |
fun_rel ~> rel_fun
|
blanchet@55942
|
2567 |
set_rel ~> rel_set
|
blanchet@55942
|
2568 |
filter_rel ~> rel_filter
|
wenzelm@57452
|
2569 |
fset_rel ~> rel_fset (in "src/HOL/Library/FSet.thy")
|
wenzelm@57452
|
2570 |
cset_rel ~> rel_cset (in "src/HOL/Library/Countable_Set_Type.thy")
|
wenzelm@57452
|
2571 |
vset ~> rel_vset (in "src/HOL/Library/Quotient_Set.thy")
|
wenzelm@57452
|
2572 |
|
wenzelm@57452
|
2573 |
INCOMPATIBILITY.
|
wenzelm@57452
|
2574 |
|
kuncar@57826
|
2575 |
* Lifting and Transfer:
|
kuncar@57826
|
2576 |
- a type variable as a raw type is supported
|
kuncar@57826
|
2577 |
- stronger reflexivity prover
|
kuncar@57826
|
2578 |
- rep_eq is always generated by lift_definition
|
wenzelm@57856
|
2579 |
- setup for Lifting/Transfer is now automated for BNFs
|
kuncar@57826
|
2580 |
+ holds for BNFs that do not contain a dead variable
|
wenzelm@57856
|
2581 |
+ relator_eq, relator_mono, relator_distr, relator_domain,
|
kuncar@57826
|
2582 |
relator_eq_onp, quot_map, transfer rules for bi_unique, bi_total,
|
kuncar@57826
|
2583 |
right_unique, right_total, left_unique, left_total are proved
|
kuncar@57826
|
2584 |
automatically
|
kuncar@57826
|
2585 |
+ definition of a predicator is generated automatically
|
kuncar@57826
|
2586 |
+ simplification rules for a predicator definition are proved
|
kuncar@57826
|
2587 |
automatically for datatypes
|
kuncar@57826
|
2588 |
- consolidation of the setup of Lifting/Transfer
|
wenzelm@57856
|
2589 |
+ property that a relator preservers reflexivity is not needed any
|
kuncar@57826
|
2590 |
more
|
kuncar@57826
|
2591 |
Minor INCOMPATIBILITY.
|
wenzelm@57856
|
2592 |
+ left_total and left_unique rules are now transfer rules
|
kuncar@57826
|
2593 |
(reflexivity_rule attribute not needed anymore)
|
kuncar@57826
|
2594 |
INCOMPATIBILITY.
|
wenzelm@57856
|
2595 |
+ Domainp does not have to be a separate assumption in
|
kuncar@57826
|
2596 |
relator_domain theorems (=> more natural statement)
|
kuncar@57826
|
2597 |
INCOMPATIBILITY.
|
kuncar@57826
|
2598 |
- registration of code equations is more robust
|
kuncar@57826
|
2599 |
Potential INCOMPATIBILITY.
|
kuncar@57826
|
2600 |
- respectfulness proof obligation is preprocessed to a more readable
|
kuncar@57826
|
2601 |
form
|
kuncar@57826
|
2602 |
Potential INCOMPATIBILITY.
|
kuncar@57826
|
2603 |
- eq_onp is always unfolded in respectfulness proof obligation
|
kuncar@57826
|
2604 |
Potential INCOMPATIBILITY.
|
wenzelm@57856
|
2605 |
- unregister lifting setup for Code_Numeral.integer and
|
kuncar@57826
|
2606 |
Code_Numeral.natural
|
kuncar@57826
|
2607 |
Potential INCOMPATIBILITY.
|
kuncar@57826
|
2608 |
- Lifting.invariant -> eq_onp
|
kuncar@57826
|
2609 |
INCOMPATIBILITY.
|
wenzelm@57856
|
2610 |
|
wenzelm@57508
|
2611 |
* New internal SAT solver "cdclite" that produces models and proof
|
wenzelm@57508
|
2612 |
traces. This solver replaces the internal SAT solvers "enumerate" and
|
wenzelm@57508
|
2613 |
"dpll". Applications that explicitly used one of these two SAT
|
wenzelm@57508
|
2614 |
solvers should use "cdclite" instead. In addition, "cdclite" is now
|
wenzelm@57508
|
2615 |
the default SAT solver for the "sat" and "satx" proof methods and
|
wenzelm@57508
|
2616 |
corresponding tactics; the old default can be restored using "declare
|
wenzelm@57508
|
2617 |
[[sat_solver = zchaff_with_proofs]]". Minor INCOMPATIBILITY.
|
wenzelm@57508
|
2618 |
|
wenzelm@57508
|
2619 |
* SMT module: A new version of the SMT module, temporarily called
|
wenzelm@57508
|
2620 |
"SMT2", uses SMT-LIB 2 and supports recent versions of Z3 (e.g.,
|
wenzelm@57508
|
2621 |
4.3). The new proof method is called "smt2". CVC3 and CVC4 are also
|
wenzelm@57508
|
2622 |
supported as oracles. Yices is no longer supported, because no version
|
wenzelm@57508
|
2623 |
of the solver can handle both SMT-LIB 2 and quantifiers.
|
wenzelm@57508
|
2624 |
|
wenzelm@57508
|
2625 |
* Activation of Z3 now works via "z3_non_commercial" system option
|
wenzelm@57508
|
2626 |
(without requiring restart), instead of former settings variable
|
wenzelm@57508
|
2627 |
"Z3_NON_COMMERCIAL". The option can be edited in Isabelle/jEdit menu
|
wenzelm@57508
|
2628 |
Plugin Options / Isabelle / General.
|
wenzelm@57508
|
2629 |
|
wenzelm@57508
|
2630 |
* Sledgehammer:
|
wenzelm@57508
|
2631 |
- Z3 can now produce Isar proofs.
|
wenzelm@57508
|
2632 |
- MaSh overhaul:
|
blanchet@57532
|
2633 |
. New SML-based learning algorithms eliminate the dependency on
|
wenzelm@57508
|
2634 |
Python and increase performance and reliability.
|
wenzelm@57508
|
2635 |
. MaSh and MeSh are now used by default together with the
|
wenzelm@57508
|
2636 |
traditional MePo (Meng-Paulson) relevance filter. To disable
|
wenzelm@57508
|
2637 |
MaSh, set the "MaSh" system option in Isabelle/jEdit Plugin
|
wenzelm@57508
|
2638 |
Options / Isabelle / General to "none".
|
wenzelm@57508
|
2639 |
- New option:
|
wenzelm@57508
|
2640 |
smt_proofs
|
wenzelm@57508
|
2641 |
- Renamed options:
|
wenzelm@57508
|
2642 |
isar_compress ~> compress
|
wenzelm@57508
|
2643 |
isar_try0 ~> try0
|
wenzelm@57508
|
2644 |
|
wenzelm@57508
|
2645 |
INCOMPATIBILITY.
|
wenzelm@57508
|
2646 |
|
wenzelm@57508
|
2647 |
* Removed solvers remote_cvc3 and remote_z3. Use cvc3 and z3 instead.
|
wenzelm@57508
|
2648 |
|
wenzelm@57508
|
2649 |
* Nitpick:
|
wenzelm@57508
|
2650 |
- Fixed soundness bug whereby mutually recursive datatypes could
|
wenzelm@57508
|
2651 |
take infinite values.
|
wenzelm@57508
|
2652 |
- Fixed soundness bug with low-level number functions such as
|
wenzelm@57508
|
2653 |
"Abs_Integ" and "Rep_Integ".
|
wenzelm@57508
|
2654 |
- Removed "std" option.
|
wenzelm@57508
|
2655 |
- Renamed "show_datatypes" to "show_types" and "hide_datatypes" to
|
wenzelm@57508
|
2656 |
"hide_types".
|
wenzelm@57508
|
2657 |
|
wenzelm@57508
|
2658 |
* Metis: Removed legacy proof method 'metisFT'. Use 'metis
|
wenzelm@57508
|
2659 |
(full_types)' instead. INCOMPATIBILITY.
|
wenzelm@57508
|
2660 |
|
wenzelm@57508
|
2661 |
* Try0: Added 'algebra' and 'meson' to the set of proof methods.
|
wenzelm@57508
|
2662 |
|
wenzelm@57508
|
2663 |
* Adjustion of INF and SUP operations:
|
wenzelm@57508
|
2664 |
- Elongated constants INFI and SUPR to INFIMUM and SUPREMUM.
|
wenzelm@57508
|
2665 |
- Consolidated theorem names containing INFI and SUPR: have INF and
|
wenzelm@57508
|
2666 |
SUP instead uniformly.
|
wenzelm@57508
|
2667 |
- More aggressive normalization of expressions involving INF and Inf
|
wenzelm@57508
|
2668 |
or SUP and Sup.
|
wenzelm@57508
|
2669 |
- INF_image and SUP_image do not unfold composition.
|
wenzelm@57508
|
2670 |
- Dropped facts INF_comp, SUP_comp.
|
wenzelm@57508
|
2671 |
- Default congruence rules strong_INF_cong and strong_SUP_cong, with
|
wenzelm@57508
|
2672 |
simplifier implication in premises. Generalize and replace former
|
wenzelm@57508
|
2673 |
INT_cong, SUP_cong
|
wenzelm@57508
|
2674 |
|
wenzelm@57508
|
2675 |
INCOMPATIBILITY.
|
wenzelm@57508
|
2676 |
|
wenzelm@57508
|
2677 |
* SUP and INF generalized to conditionally_complete_lattice.
|
wenzelm@57508
|
2678 |
|
wenzelm@57508
|
2679 |
* Swapped orientation of facts image_comp and vimage_comp:
|
wenzelm@57508
|
2680 |
|
wenzelm@57508
|
2681 |
image_compose ~> image_comp [symmetric]
|
wenzelm@57508
|
2682 |
image_comp ~> image_comp [symmetric]
|
wenzelm@57508
|
2683 |
vimage_compose ~> vimage_comp [symmetric]
|
wenzelm@57508
|
2684 |
vimage_comp ~> vimage_comp [symmetric]
|
wenzelm@57508
|
2685 |
|
wenzelm@57508
|
2686 |
INCOMPATIBILITY.
|
wenzelm@57508
|
2687 |
|
wenzelm@57504
|
2688 |
* Theory reorganization: split of Big_Operators.thy into
|
wenzelm@57504
|
2689 |
Groups_Big.thy and Lattices_Big.thy.
|
blanchet@55098
|
2690 |
|
haftmann@57418
|
2691 |
* Consolidated some facts about big group operators:
|
haftmann@57418
|
2692 |
|
haftmann@57418
|
2693 |
setsum_0' ~> setsum.neutral
|
haftmann@57418
|
2694 |
setsum_0 ~> setsum.neutral_const
|
haftmann@57418
|
2695 |
setsum_addf ~> setsum.distrib
|
haftmann@57418
|
2696 |
setsum_cartesian_product ~> setsum.cartesian_product
|
haftmann@57418
|
2697 |
setsum_cases ~> setsum.If_cases
|
haftmann@57418
|
2698 |
setsum_commute ~> setsum.commute
|
haftmann@57418
|
2699 |
setsum_cong ~> setsum.cong
|
haftmann@57418
|
2700 |
setsum_delta ~> setsum.delta
|
haftmann@57418
|
2701 |
setsum_delta' ~> setsum.delta'
|
haftmann@57418
|
2702 |
setsum_diff1' ~> setsum.remove
|
haftmann@57418
|
2703 |
setsum_empty ~> setsum.empty
|
haftmann@57418
|
2704 |
setsum_infinite ~> setsum.infinite
|
haftmann@57418
|
2705 |
setsum_insert ~> setsum.insert
|
haftmann@57418
|
2706 |
setsum_inter_restrict'' ~> setsum.inter_filter
|
haftmann@57418
|
2707 |
setsum_mono_zero_cong_left ~> setsum.mono_neutral_cong_left
|
haftmann@57418
|
2708 |
setsum_mono_zero_cong_right ~> setsum.mono_neutral_cong_right
|
haftmann@57418
|
2709 |
setsum_mono_zero_left ~> setsum.mono_neutral_left
|
haftmann@57418
|
2710 |
setsum_mono_zero_right ~> setsum.mono_neutral_right
|
haftmann@57418
|
2711 |
setsum_reindex ~> setsum.reindex
|
haftmann@57418
|
2712 |
setsum_reindex_cong ~> setsum.reindex_cong
|
haftmann@57418
|
2713 |
setsum_reindex_nonzero ~> setsum.reindex_nontrivial
|
haftmann@57418
|
2714 |
setsum_restrict_set ~> setsum.inter_restrict
|
haftmann@57418
|
2715 |
setsum_Plus ~> setsum.Plus
|
haftmann@57418
|
2716 |
setsum_setsum_restrict ~> setsum.commute_restrict
|
haftmann@57418
|
2717 |
setsum_Sigma ~> setsum.Sigma
|
haftmann@57418
|
2718 |
setsum_subset_diff ~> setsum.subset_diff
|
haftmann@57418
|
2719 |
setsum_Un_disjoint ~> setsum.union_disjoint
|
haftmann@57418
|
2720 |
setsum_UN_disjoint ~> setsum.UNION_disjoint
|
haftmann@57418
|
2721 |
setsum_Un_Int ~> setsum.union_inter
|
haftmann@57418
|
2722 |
setsum_Union_disjoint ~> setsum.Union_disjoint
|
haftmann@57418
|
2723 |
setsum_UNION_zero ~> setsum.Union_comp
|
haftmann@57418
|
2724 |
setsum_Un_zero ~> setsum.union_inter_neutral
|
haftmann@57418
|
2725 |
strong_setprod_cong ~> setprod.strong_cong
|
haftmann@57418
|
2726 |
strong_setsum_cong ~> setsum.strong_cong
|
haftmann@57418
|
2727 |
setprod_1' ~> setprod.neutral
|
haftmann@57418
|
2728 |
setprod_1 ~> setprod.neutral_const
|
haftmann@57418
|
2729 |
setprod_cartesian_product ~> setprod.cartesian_product
|
haftmann@57418
|
2730 |
setprod_cong ~> setprod.cong
|
haftmann@57418
|
2731 |
setprod_delta ~> setprod.delta
|
haftmann@57418
|
2732 |
setprod_delta' ~> setprod.delta'
|
haftmann@57418
|
2733 |
setprod_empty ~> setprod.empty
|
haftmann@57418
|
2734 |
setprod_infinite ~> setprod.infinite
|
haftmann@57418
|
|