wenzelm@5363
|
1 |
Isabelle NEWS -- history user-relevant changes
|
wenzelm@5363
|
2 |
==============================================
|
wenzelm@2553
|
3 |
|
wenzelm@17754
|
4 |
New in this Isabelle release
|
wenzelm@17754
|
5 |
----------------------------
|
wenzelm@17754
|
6 |
|
wenzelm@17754
|
7 |
*** General ***
|
wenzelm@17754
|
8 |
|
wenzelm@17918
|
9 |
* Theory syntax: the header format ``theory A = B + C:'' has been
|
wenzelm@17918
|
10 |
discontinued in favour of ``theory A imports B C begin''. Use isatool
|
wenzelm@17918
|
11 |
fixheaders to convert existing theory files. INCOMPATIBILITY.
|
wenzelm@17918
|
12 |
|
wenzelm@17918
|
13 |
* Theory syntax: the old non-Isar theory file format has been
|
wenzelm@17918
|
14 |
discontinued altogether. Note that ML proof scripts may still be used
|
wenzelm@17918
|
15 |
with Isar theories; migration is usually quite simple with the ML
|
wenzelm@17918
|
16 |
function use_legacy_bindings. INCOMPATIBILITY.
|
wenzelm@17918
|
17 |
|
wenzelm@19814
|
18 |
* Theory syntax: some popular names (e.g. "class", "if") are now
|
wenzelm@19814
|
19 |
keywords. INCOMPATIBILITY, use double quotes.
|
wenzelm@19814
|
20 |
|
wenzelm@17981
|
21 |
* Legacy goal package: reduced interface to the bare minimum required
|
wenzelm@17981
|
22 |
to keep existing proof scripts running. Most other user-level
|
wenzelm@17981
|
23 |
functions are now part of the OldGoals structure, which is *not* open
|
wenzelm@17981
|
24 |
by default (consider isatool expandshort before open OldGoals).
|
wenzelm@17981
|
25 |
Removed top_sg, prin, printyp, pprint_term/typ altogether, because
|
wenzelm@17981
|
26 |
these tend to cause confusion about the actual goal (!) context being
|
wenzelm@17981
|
27 |
used here, which is not necessarily the same as the_context().
|
wenzelm@17918
|
28 |
|
wenzelm@17754
|
29 |
* Command 'find_theorems': support "*" wildcard in "name:" criterion.
|
wenzelm@17754
|
30 |
|
wenzelm@20370
|
31 |
* The ``prems limit'' option (cf. ProofContext.prems_limit) is now -1
|
wenzelm@20370
|
32 |
by default, which means that "prems" (and also "fixed variables") are
|
wenzelm@20370
|
33 |
suppressed from proof state output. Note that the ProofGeneral
|
wenzelm@20370
|
34 |
settings mechanism allows to change and save options persistently, but
|
wenzelm@20370
|
35 |
older versions of Isabelle will fail to start up if a negative prems
|
wenzelm@20370
|
36 |
limit is imposed.
|
wenzelm@20370
|
37 |
|
wenzelm@17754
|
38 |
|
wenzelm@17865
|
39 |
*** Document preparation ***
|
wenzelm@17865
|
40 |
|
wenzelm@17869
|
41 |
* Added antiquotations @{ML_type text} and @{ML_struct text} which
|
wenzelm@17869
|
42 |
check the given source text as ML type/structure, printing verbatim.
|
wenzelm@17865
|
43 |
|
wenzelm@17865
|
44 |
|
wenzelm@17779
|
45 |
*** Pure ***
|
wenzelm@17779
|
46 |
|
haftmann@20188
|
47 |
* class_package.ML offers a combination of axclasses and locales to achieve
|
haftmann@20188
|
48 |
Haskell-like type classes in Isabelle. See HOL/ex/Classpackage.thy for examples.
|
haftmann@20188
|
49 |
|
haftmann@20188
|
50 |
* Yet another code generator framework allows to generate executable code
|
haftmann@20453
|
51 |
for ML and Haskell (including "class"es). A short usage sketch:
|
haftmann@20188
|
52 |
|
haftmann@20188
|
53 |
internal compilation:
|
haftmann@20453
|
54 |
code_gen <list of constants (term syntax)> (SML -)
|
haftmann@20453
|
55 |
writing SML code to a file:
|
haftmann@20453
|
56 |
code_gen <list of constants (term syntax)> (SML <filename>)
|
haftmann@20188
|
57 |
writing Haskell code to a bunch of files:
|
haftmann@20453
|
58 |
code_gen <list of constants (term syntax)> (Haskell <filename>)
|
haftmann@20453
|
59 |
|
haftmann@20453
|
60 |
Reasonable default setup of framework in HOL/Main.
|
haftmann@20453
|
61 |
|
haftmann@20453
|
62 |
See HOL/ex/Code*.thy for examples.
|
haftmann@20453
|
63 |
|
haftmann@20453
|
64 |
Theorem attributs for selecting and transforming function equations theorems:
|
haftmann@20453
|
65 |
|
haftmann@20453
|
66 |
[code fun]: select a theorem as function equation for a specific constant
|
haftmann@20453
|
67 |
[code nofun]: deselect a theorem as function equation for a specific constant
|
haftmann@20453
|
68 |
[code inline]: select an equation theorem for unfolding (inlining) in place
|
haftmann@20453
|
69 |
[code noinline]: deselect an equation theorem for unfolding (inlining) in place
|
haftmann@20453
|
70 |
|
haftmann@20453
|
71 |
User-defined serializations (target in {SML, Haskell}):
|
haftmann@20453
|
72 |
|
haftmann@20453
|
73 |
code_const <and-list of constants (term syntax)>
|
haftmann@20453
|
74 |
{(target) <and-list of const target syntax>}+
|
haftmann@20453
|
75 |
|
haftmann@20453
|
76 |
code_type <and-list of type constructors>
|
haftmann@20453
|
77 |
{(target) <and-list of type target syntax>}+
|
haftmann@20453
|
78 |
|
haftmann@20453
|
79 |
code_instance <and-list of instances>
|
haftmann@20453
|
80 |
{(target)}+
|
haftmann@20453
|
81 |
where instance ::= <type constructor> :: <class>
|
haftmann@20453
|
82 |
|
haftmann@20453
|
83 |
code_class <and_list of classes>
|
haftmann@20453
|
84 |
{(target) <and-list of class target syntax>}+
|
haftmann@20453
|
85 |
where class target syntax ::= <class name> {where {<classop> == <target syntax>}+}?
|
haftmann@20453
|
86 |
|
haftmann@20453
|
87 |
For code_instance and code_class, target SML is silently ignored.
|
haftmann@20453
|
88 |
|
haftmann@20453
|
89 |
See HOL theories and HOL/ex/Code*.thy for usage examples.
|
haftmann@20188
|
90 |
|
wenzelm@19254
|
91 |
* Command 'no_translations' removes translation rules from theory
|
wenzelm@19254
|
92 |
syntax.
|
wenzelm@19254
|
93 |
|
wenzelm@19625
|
94 |
* Overloaded definitions are now actually checked for acyclic
|
wenzelm@19714
|
95 |
dependencies. The overloading scheme is slightly more general than
|
wenzelm@19714
|
96 |
that of Haskell98, although Isabelle does not demand an exact
|
wenzelm@19714
|
97 |
correspondence to type class and instance declarations.
|
wenzelm@19714
|
98 |
INCOMPATIBILITY, use ``defs (unchecked overloaded)'' to admit more
|
wenzelm@19714
|
99 |
exotic versions of overloading -- at the discretion of the user!
|
wenzelm@19711
|
100 |
|
wenzelm@19711
|
101 |
Polymorphic constants are represented via type arguments, i.e. the
|
wenzelm@19711
|
102 |
instantiation that matches an instance against the most general
|
wenzelm@19711
|
103 |
declaration given in the signature. For example, with the declaration
|
wenzelm@19711
|
104 |
c :: 'a => 'a => 'a, an instance c :: nat => nat => nat is represented
|
wenzelm@19711
|
105 |
as c(nat). Overloading is essentially simultaneous structural
|
wenzelm@19711
|
106 |
recursion over such type arguments. Incomplete specification patterns
|
wenzelm@19714
|
107 |
impose global constraints on all occurrences, e.g. c('a * 'a) on the
|
wenzelm@19715
|
108 |
LHS means that more general c('a * 'b) will be disallowed on any RHS.
|
wenzelm@19714
|
109 |
Command 'print_theory' outputs the normalized system of recursive
|
wenzelm@19714
|
110 |
equations, see section "definitions".
|
wenzelm@19625
|
111 |
|
wenzelm@17865
|
112 |
* Isar: improper proof element 'guess' is like 'obtain', but derives
|
wenzelm@17865
|
113 |
the obtained context from the course of reasoning! For example:
|
wenzelm@17865
|
114 |
|
wenzelm@17865
|
115 |
assume "EX x y. A x & B y" -- "any previous fact"
|
wenzelm@17865
|
116 |
then guess x and y by clarify
|
wenzelm@17865
|
117 |
|
wenzelm@17865
|
118 |
This technique is potentially adventurous, depending on the facts and
|
wenzelm@17865
|
119 |
proof tools being involved here.
|
wenzelm@17865
|
120 |
|
wenzelm@18020
|
121 |
* Isar: known facts from the proof context may be specified as literal
|
wenzelm@18020
|
122 |
propositions, using ASCII back-quote syntax. This works wherever
|
wenzelm@18020
|
123 |
named facts used to be allowed so far, in proof commands, proof
|
wenzelm@18020
|
124 |
methods, attributes etc. Literal facts are retrieved from the context
|
wenzelm@18020
|
125 |
according to unification of type and term parameters. For example,
|
wenzelm@18020
|
126 |
provided that "A" and "A ==> B" and "!!x. P x ==> Q x" are known
|
wenzelm@18020
|
127 |
theorems in the current context, then these are valid literal facts:
|
wenzelm@18020
|
128 |
`A` and `A ==> B` and `!!x. P x ==> Q x" as well as `P a ==> Q a` etc.
|
wenzelm@18020
|
129 |
|
wenzelm@18020
|
130 |
There is also a proof method "fact" which does the same composition
|
wenzelm@18044
|
131 |
for explicit goal states, e.g. the following proof texts coincide with
|
wenzelm@18044
|
132 |
certain special cases of literal facts:
|
wenzelm@18020
|
133 |
|
wenzelm@18020
|
134 |
have "A" by fact == note `A`
|
wenzelm@18020
|
135 |
have "A ==> B" by fact == note `A ==> B`
|
wenzelm@18020
|
136 |
have "!!x. P x ==> Q x" by fact == note `!!x. P x ==> Q x`
|
wenzelm@18020
|
137 |
have "P a ==> Q a" by fact == note `P a ==> Q a`
|
wenzelm@18020
|
138 |
|
wenzelm@20118
|
139 |
* Isar: ":" (colon) is no longer a symbolic identifier character in
|
wenzelm@20118
|
140 |
outer syntax. Thus symbolic identifiers may be used without
|
wenzelm@20118
|
141 |
additional white space in declarations like this: ``assume *: A''.
|
wenzelm@20118
|
142 |
|
wenzelm@20013
|
143 |
* Isar: 'print_facts' prints all local facts of the current context,
|
wenzelm@20013
|
144 |
both named and unnamed ones.
|
wenzelm@20013
|
145 |
|
wenzelm@18308
|
146 |
* Isar: 'def' now admits simultaneous definitions, e.g.:
|
wenzelm@18308
|
147 |
|
wenzelm@18308
|
148 |
def x == "t" and y == "u"
|
wenzelm@18308
|
149 |
|
wenzelm@18540
|
150 |
* Isar: added command 'unfolding', which is structurally similar to
|
wenzelm@18540
|
151 |
'using', but affects both the goal state and facts by unfolding given
|
wenzelm@18815
|
152 |
rewrite rules. Thus many occurrences of the 'unfold' method or
|
wenzelm@18540
|
153 |
'unfolded' attribute may be replaced by first-class proof text.
|
wenzelm@18540
|
154 |
|
wenzelm@18815
|
155 |
* Isar: methods 'unfold' / 'fold', attributes 'unfolded' / 'folded',
|
wenzelm@18815
|
156 |
and command 'unfolding' now all support object-level equalities
|
wenzelm@18815
|
157 |
(potentially conditional). The underlying notion of rewrite rule is
|
wenzelm@18815
|
158 |
analogous to the 'rule_format' attribute, but *not* that of the
|
wenzelm@18815
|
159 |
Simplifier (which is usually more generous).
|
wenzelm@18815
|
160 |
|
wenzelm@19220
|
161 |
* Isar: the goal restriction operator [N] (default N = 1) evaluates a
|
wenzelm@19220
|
162 |
method expression within a sandbox consisting of the first N
|
wenzelm@19240
|
163 |
sub-goals, which need to exist. For example, ``simp_all [3]''
|
wenzelm@19240
|
164 |
simplifies the first three sub-goals, while (rule foo, simp_all)[]
|
wenzelm@19240
|
165 |
simplifies all new goals that emerge from applying rule foo to the
|
wenzelm@19240
|
166 |
originally first one.
|
wenzelm@19220
|
167 |
|
wenzelm@19814
|
168 |
* Isar: schematic goals are no longer restricted to higher-order
|
wenzelm@19814
|
169 |
patterns; e.g. ``lemma "?P(?x)" by (rule TrueI)'' now works as
|
wenzelm@19814
|
170 |
expected.
|
wenzelm@19814
|
171 |
|
wenzelm@18901
|
172 |
* Isar: the conclusion of a long theorem statement is now either
|
wenzelm@18901
|
173 |
'shows' (a simultaneous conjunction, as before), or 'obtains'
|
wenzelm@18901
|
174 |
(essentially a disjunction of cases with local parameters and
|
wenzelm@18901
|
175 |
assumptions). The latter allows to express general elimination rules
|
wenzelm@18910
|
176 |
adequately; in this notation common elimination rules look like this:
|
wenzelm@18901
|
177 |
|
wenzelm@18901
|
178 |
lemma exE: -- "EX x. P x ==> (!!x. P x ==> thesis) ==> thesis"
|
wenzelm@18901
|
179 |
assumes "EX x. P x"
|
wenzelm@18901
|
180 |
obtains x where "P x"
|
wenzelm@18901
|
181 |
|
wenzelm@18901
|
182 |
lemma conjE: -- "A & B ==> (A ==> B ==> thesis) ==> thesis"
|
wenzelm@18901
|
183 |
assumes "A & B"
|
wenzelm@18901
|
184 |
obtains A and B
|
wenzelm@18901
|
185 |
|
wenzelm@18901
|
186 |
lemma disjE: -- "A | B ==> (A ==> thesis) ==> (B ==> thesis) ==> thesis"
|
wenzelm@18901
|
187 |
assumes "A | B"
|
wenzelm@18901
|
188 |
obtains
|
wenzelm@18901
|
189 |
A
|
wenzelm@18901
|
190 |
| B
|
wenzelm@18901
|
191 |
|
wenzelm@18910
|
192 |
The subsequent classical rules even refer to the formal "thesis"
|
wenzelm@18901
|
193 |
explicitly:
|
wenzelm@18901
|
194 |
|
wenzelm@18901
|
195 |
lemma classical: -- "(~ thesis ==> thesis) ==> thesis"
|
wenzelm@18901
|
196 |
obtains "~ thesis"
|
wenzelm@18901
|
197 |
|
wenzelm@18910
|
198 |
lemma Peirce's_Law: -- "((thesis ==> something) ==> thesis) ==> thesis"
|
wenzelm@18910
|
199 |
obtains "thesis ==> something"
|
wenzelm@18901
|
200 |
|
wenzelm@18901
|
201 |
The actual proof of an 'obtains' statement is analogous to that of the
|
wenzelm@18910
|
202 |
Isar proof element 'obtain', only that there may be several cases.
|
wenzelm@18910
|
203 |
Optional case names may be specified in parentheses; these will be
|
wenzelm@18910
|
204 |
available both in the present proof and as annotations in the
|
wenzelm@18910
|
205 |
resulting rule, for later use with the 'cases' method (cf. attribute
|
wenzelm@18910
|
206 |
case_names).
|
wenzelm@18901
|
207 |
|
wenzelm@19263
|
208 |
* Isar: 'print_statement' prints theorems from the current theory or
|
wenzelm@19263
|
209 |
proof context in long statement form, according to the syntax of a
|
wenzelm@19263
|
210 |
top-level lemma.
|
wenzelm@19263
|
211 |
|
wenzelm@18901
|
212 |
* Isar: 'obtain' takes an optional case name for the local context
|
wenzelm@18901
|
213 |
introduction rule (default "that").
|
wenzelm@18901
|
214 |
|
wenzelm@19587
|
215 |
* Isar: removed obsolete 'concl is' patterns. INCOMPATIBILITY, use
|
wenzelm@19587
|
216 |
explicit (is "_ ==> ?foo") in the rare cases where this still happens
|
wenzelm@19587
|
217 |
to occur.
|
wenzelm@19587
|
218 |
|
wenzelm@19682
|
219 |
* Pure: syntax "CONST name" produces a fully internalized constant
|
wenzelm@19682
|
220 |
according to the current context. This is particularly useful for
|
wenzelm@19682
|
221 |
syntax translations that should refer to internal constant
|
wenzelm@19682
|
222 |
representations independently of name spaces.
|
wenzelm@19682
|
223 |
|
wenzelm@19682
|
224 |
* Isar/locales: 'const_syntax' provides a robust interface to the
|
wenzelm@19682
|
225 |
'syntax' primitive that also works in a locale context. Type
|
wenzelm@19682
|
226 |
declaration and internal syntactic representation of given constants
|
wenzelm@19682
|
227 |
retrieved from the context.
|
wenzelm@19682
|
228 |
|
wenzelm@19665
|
229 |
* Isar/locales: new derived specification elements 'axiomatization',
|
wenzelm@19665
|
230 |
'definition', 'abbreviation', which support type-inference, admit
|
wenzelm@19083
|
231 |
object-level specifications (equality, equivalence). See also the
|
wenzelm@19083
|
232 |
isar-ref manual. Examples:
|
wenzelm@19081
|
233 |
|
wenzelm@19665
|
234 |
axiomatization
|
wenzelm@19665
|
235 |
eq (infix "===" 50)
|
wenzelm@19665
|
236 |
where eq_refl: "x === x" and eq_subst: "x === y ==> P x ==> P y"
|
wenzelm@19665
|
237 |
|
wenzelm@19081
|
238 |
definition
|
wenzelm@19081
|
239 |
"f x y = x + y + 1"
|
wenzelm@19081
|
240 |
"g x = f x x"
|
wenzelm@19081
|
241 |
|
wenzelm@19363
|
242 |
abbreviation
|
wenzelm@19081
|
243 |
neq (infix "=!=" 50)
|
wenzelm@19363
|
244 |
"x =!= y == ~ (x === y)"
|
wenzelm@19081
|
245 |
|
wenzelm@19083
|
246 |
These specifications may be also used in a locale context. Then the
|
wenzelm@19083
|
247 |
constants being introduced depend on certain fixed parameters, and the
|
wenzelm@19083
|
248 |
constant name is qualified by the locale base name. An internal
|
wenzelm@19083
|
249 |
abbreviation takes care for convenient input and output, making the
|
wenzelm@19088
|
250 |
parameters implicit and using the original short name. See also
|
wenzelm@19083
|
251 |
HOL/ex/Abstract_NAT.thy for an example of deriving polymorphic
|
wenzelm@19083
|
252 |
entities from a monomorphic theory.
|
wenzelm@19083
|
253 |
|
wenzelm@19083
|
254 |
Presently, abbreviations are only available 'in' a target locale, but
|
wenzelm@19363
|
255 |
not inherited by general import expressions. Also note that
|
wenzelm@19363
|
256 |
'abbreviation' may be used as a type-safe replacement for 'syntax' +
|
wenzelm@19363
|
257 |
'translations' in common applications.
|
wenzelm@19084
|
258 |
|
wenzelm@19682
|
259 |
Concrete syntax is attached to specified constants in internal form,
|
wenzelm@19682
|
260 |
independently of name spaces. The parse tree representation is
|
wenzelm@19682
|
261 |
slightly different -- use 'const_syntax' instead of raw 'syntax', and
|
wenzelm@19682
|
262 |
'translations' with explicit "CONST" markup to accommodate this.
|
wenzelm@19665
|
263 |
|
ballarin@19783
|
264 |
* Isar/locales: improved parameter handling:
|
ballarin@19783
|
265 |
- use of locales "var" and "struct" no longer necessary;
|
ballarin@19783
|
266 |
- parameter renamings are no longer required to be injective.
|
ballarin@19783
|
267 |
This enables, for example, to define a locale for endomorphisms thus:
|
ballarin@19783
|
268 |
locale endom = homom mult mult h.
|
ballarin@19783
|
269 |
|
ballarin@19931
|
270 |
* Isar/locales: changed the way locales with predicates are defined.
|
ballarin@19931
|
271 |
Instead of accumulating the specification, the imported expression is
|
ballarin@19931
|
272 |
now an interpretation.
|
ballarin@19931
|
273 |
INCOMPATIBILITY: different normal form of locale expressions.
|
ballarin@19931
|
274 |
In particular, in interpretations of locales with predicates,
|
ballarin@19931
|
275 |
goals repesenting already interpreted fragments are not removed
|
ballarin@19984
|
276 |
automatically. Use methods `intro_locales' and `unfold_locales'; see below.
|
ballarin@19984
|
277 |
|
ballarin@19984
|
278 |
* Isar/locales: new methods `intro_locales' and `unfold_locales' provide
|
ballarin@19984
|
279 |
backward reasoning on locales predicates. The methods are aware of
|
ballarin@19984
|
280 |
interpretations and discharge corresponding goals. `intro_locales' is
|
ballarin@19984
|
281 |
less aggressive then `unfold_locales' and does not unfold predicates to
|
ballarin@19984
|
282 |
assumptions.
|
ballarin@19931
|
283 |
|
ballarin@19931
|
284 |
* Isar/locales: the order in which locale fragments are accumulated
|
ballarin@19931
|
285 |
has changed. This enables to override declarations from fragments
|
ballarin@19931
|
286 |
due to interpretations -- for example, unwanted simp rules.
|
ballarin@19931
|
287 |
|
wenzelm@18233
|
288 |
* Provers/induct: improved internal context management to support
|
wenzelm@18233
|
289 |
local fixes and defines on-the-fly. Thus explicit meta-level
|
wenzelm@18233
|
290 |
connectives !! and ==> are rarely required anymore in inductive goals
|
wenzelm@18233
|
291 |
(using object-logic connectives for this purpose has been long
|
wenzelm@18233
|
292 |
obsolete anyway). The subsequent proof patterns illustrate advanced
|
wenzelm@18233
|
293 |
techniques of natural induction; general datatypes and inductive sets
|
wenzelm@18267
|
294 |
work analogously (see also src/HOL/Lambda for realistic examples).
|
wenzelm@18267
|
295 |
|
wenzelm@18267
|
296 |
(1) This is how to ``strengthen'' an inductive goal wrt. certain
|
wenzelm@18239
|
297 |
parameters:
|
wenzelm@18233
|
298 |
|
wenzelm@18233
|
299 |
lemma
|
wenzelm@18233
|
300 |
fixes n :: nat and x :: 'a
|
wenzelm@18233
|
301 |
assumes a: "A n x"
|
wenzelm@18233
|
302 |
shows "P n x"
|
wenzelm@18233
|
303 |
using a -- {* make induct insert fact a *}
|
wenzelm@20503
|
304 |
proof (induct n arbitrary: x) -- {* generalize goal to "!!x. A n x ==> P n x" *}
|
wenzelm@18248
|
305 |
case 0
|
wenzelm@18233
|
306 |
show ?case sorry
|
wenzelm@18233
|
307 |
next
|
wenzelm@18248
|
308 |
case (Suc n)
|
wenzelm@18239
|
309 |
note `!!x. A n x ==> P n x` -- {* induction hypothesis, according to induction rule *}
|
wenzelm@18239
|
310 |
note `A (Suc n) x` -- {* induction premise, stemming from fact a *}
|
wenzelm@18233
|
311 |
show ?case sorry
|
wenzelm@18233
|
312 |
qed
|
wenzelm@18233
|
313 |
|
wenzelm@18267
|
314 |
(2) This is how to perform induction over ``expressions of a certain
|
wenzelm@18233
|
315 |
form'', using a locally defined inductive parameter n == "a x"
|
wenzelm@18239
|
316 |
together with strengthening (the latter is usually required to get
|
wenzelm@18267
|
317 |
sufficiently flexible induction hypotheses):
|
wenzelm@18233
|
318 |
|
wenzelm@18233
|
319 |
lemma
|
wenzelm@18233
|
320 |
fixes a :: "'a => nat"
|
wenzelm@18233
|
321 |
assumes a: "A (a x)"
|
wenzelm@18233
|
322 |
shows "P (a x)"
|
wenzelm@18233
|
323 |
using a
|
wenzelm@20503
|
324 |
proof (induct n == "a x" arbitrary: x)
|
wenzelm@18233
|
325 |
...
|
wenzelm@18233
|
326 |
|
wenzelm@18267
|
327 |
See also HOL/Isar_examples/Puzzle.thy for an application of the this
|
wenzelm@18267
|
328 |
particular technique.
|
wenzelm@18267
|
329 |
|
wenzelm@18901
|
330 |
(3) This is how to perform existential reasoning ('obtains' or
|
wenzelm@18901
|
331 |
'obtain') by induction, while avoiding explicit object-logic
|
wenzelm@18901
|
332 |
encodings:
|
wenzelm@18901
|
333 |
|
wenzelm@18901
|
334 |
lemma
|
wenzelm@18901
|
335 |
fixes n :: nat
|
wenzelm@18901
|
336 |
obtains x :: 'a where "P n x" and "Q n x"
|
wenzelm@20503
|
337 |
proof (induct n arbitrary: thesis)
|
wenzelm@18267
|
338 |
case 0
|
wenzelm@18267
|
339 |
obtain x where "P 0 x" and "Q 0 x" sorry
|
wenzelm@18399
|
340 |
then show thesis by (rule 0)
|
wenzelm@18267
|
341 |
next
|
wenzelm@18267
|
342 |
case (Suc n)
|
wenzelm@18267
|
343 |
obtain x where "P n x" and "Q n x" by (rule Suc.hyps)
|
wenzelm@18267
|
344 |
obtain x where "P (Suc n) x" and "Q (Suc n) x" sorry
|
wenzelm@18267
|
345 |
then show thesis by (rule Suc.prems)
|
wenzelm@18267
|
346 |
qed
|
wenzelm@18267
|
347 |
|
wenzelm@20503
|
348 |
Here the 'arbitrary: thesis' specification essentially modifies the
|
wenzelm@20503
|
349 |
scope of the formal thesis parameter, in order to the get the whole
|
wenzelm@18267
|
350 |
existence statement through the induction as expected.
|
wenzelm@18233
|
351 |
|
wenzelm@18506
|
352 |
* Provers/induct: mutual induction rules are now specified as a list
|
wenzelm@18506
|
353 |
of rule sharing the same induction cases. HOL packages usually
|
wenzelm@18506
|
354 |
provide foo_bar.inducts for mutually defined items foo and bar
|
wenzelm@18506
|
355 |
(e.g. inductive sets or datatypes). INCOMPATIBILITY, users need to
|
wenzelm@18506
|
356 |
specify mutual induction rules differently, i.e. like this:
|
wenzelm@18506
|
357 |
|
wenzelm@18506
|
358 |
(induct rule: foo_bar.inducts)
|
wenzelm@18506
|
359 |
(induct set: foo bar)
|
wenzelm@18506
|
360 |
(induct type: foo bar)
|
wenzelm@18506
|
361 |
|
wenzelm@18506
|
362 |
The ML function ProjectRule.projections turns old-style rules into the
|
wenzelm@18506
|
363 |
new format.
|
wenzelm@18506
|
364 |
|
wenzelm@18506
|
365 |
* Provers/induct: improved handling of simultaneous goals. Instead of
|
wenzelm@18506
|
366 |
introducing object-level conjunction, the statement is now split into
|
wenzelm@18506
|
367 |
several conclusions, while the corresponding symbolic cases are
|
wenzelm@18601
|
368 |
nested accordingly. INCOMPATIBILITY, proofs need to be structured
|
wenzelm@18601
|
369 |
explicitly. For example:
|
wenzelm@18480
|
370 |
|
wenzelm@18480
|
371 |
lemma
|
wenzelm@18480
|
372 |
fixes n :: nat
|
wenzelm@18480
|
373 |
shows "P n" and "Q n"
|
wenzelm@18480
|
374 |
proof (induct n)
|
wenzelm@18601
|
375 |
case 0 case 1
|
wenzelm@18480
|
376 |
show "P 0" sorry
|
wenzelm@18480
|
377 |
next
|
wenzelm@18601
|
378 |
case 0 case 2
|
wenzelm@18480
|
379 |
show "Q 0" sorry
|
wenzelm@18480
|
380 |
next
|
wenzelm@18601
|
381 |
case (Suc n) case 1
|
wenzelm@18480
|
382 |
note `P n` and `Q n`
|
wenzelm@18480
|
383 |
show "P (Suc n)" sorry
|
wenzelm@18480
|
384 |
next
|
wenzelm@18601
|
385 |
case (Suc n) case 2
|
wenzelm@18480
|
386 |
note `P n` and `Q n`
|
wenzelm@18480
|
387 |
show "Q (Suc n)" sorry
|
wenzelm@18480
|
388 |
qed
|
wenzelm@18480
|
389 |
|
wenzelm@18601
|
390 |
The split into subcases may be deferred as follows -- this is
|
wenzelm@18601
|
391 |
particularly relevant for goal statements with local premises.
|
wenzelm@18601
|
392 |
|
wenzelm@18601
|
393 |
lemma
|
wenzelm@18601
|
394 |
fixes n :: nat
|
wenzelm@18601
|
395 |
shows "A n ==> P n" and "B n ==> Q n"
|
wenzelm@18601
|
396 |
proof (induct n)
|
wenzelm@18601
|
397 |
case 0
|
wenzelm@18601
|
398 |
{
|
wenzelm@18601
|
399 |
case 1
|
wenzelm@18601
|
400 |
note `A 0`
|
wenzelm@18601
|
401 |
show "P 0" sorry
|
wenzelm@18601
|
402 |
next
|
wenzelm@18601
|
403 |
case 2
|
wenzelm@18601
|
404 |
note `B 0`
|
wenzelm@18601
|
405 |
show "Q 0" sorry
|
wenzelm@18601
|
406 |
}
|
wenzelm@18601
|
407 |
next
|
wenzelm@18601
|
408 |
case (Suc n)
|
wenzelm@18601
|
409 |
note `A n ==> P n` and `B n ==> Q n`
|
wenzelm@18601
|
410 |
{
|
wenzelm@18601
|
411 |
case 1
|
wenzelm@18601
|
412 |
note `A (Suc n)`
|
wenzelm@18601
|
413 |
show "P (Suc n)" sorry
|
wenzelm@18601
|
414 |
next
|
wenzelm@18601
|
415 |
case 2
|
wenzelm@18601
|
416 |
note `B (Suc n)`
|
wenzelm@18601
|
417 |
show "Q (Suc n)" sorry
|
wenzelm@18601
|
418 |
}
|
wenzelm@18601
|
419 |
qed
|
wenzelm@18601
|
420 |
|
wenzelm@18506
|
421 |
If simultaneous goals are to be used with mutual rules, the statement
|
wenzelm@18506
|
422 |
needs to be structured carefully as a two-level conjunction, using
|
wenzelm@18506
|
423 |
lists of propositions separated by 'and':
|
wenzelm@18506
|
424 |
|
wenzelm@18507
|
425 |
lemma
|
wenzelm@18507
|
426 |
shows "a : A ==> P1 a"
|
wenzelm@18507
|
427 |
"a : A ==> P2 a"
|
wenzelm@18507
|
428 |
and "b : B ==> Q1 b"
|
wenzelm@18507
|
429 |
"b : B ==> Q2 b"
|
wenzelm@18507
|
430 |
"b : B ==> Q3 b"
|
wenzelm@18507
|
431 |
proof (induct set: A B)
|
wenzelm@18480
|
432 |
|
wenzelm@18399
|
433 |
* Provers/induct: support coinduction as well. See
|
wenzelm@18399
|
434 |
src/HOL/Library/Coinductive_List.thy for various examples.
|
wenzelm@18399
|
435 |
|
nipkow@18674
|
436 |
* Simplifier: by default the simplifier trace only shows top level rewrites
|
nipkow@18674
|
437 |
now. That is, trace_simp_depth_limit is set to 1 by default. Thus there is
|
nipkow@18674
|
438 |
less danger of being flooded by the trace. The trace indicates where parts
|
nipkow@18674
|
439 |
have been suppressed.
|
nipkow@18674
|
440 |
|
wenzelm@18536
|
441 |
* Provers/classical: removed obsolete classical version of elim_format
|
wenzelm@18536
|
442 |
attribute; classical elim/dest rules are now treated uniformly when
|
wenzelm@18536
|
443 |
manipulating the claset.
|
wenzelm@18536
|
444 |
|
wenzelm@18694
|
445 |
* Provers/classical: stricter checks to ensure that supplied intro,
|
wenzelm@18694
|
446 |
dest and elim rules are well-formed; dest and elim rules must have at
|
wenzelm@18694
|
447 |
least one premise.
|
wenzelm@18694
|
448 |
|
wenzelm@18694
|
449 |
* Provers/classical: attributes dest/elim/intro take an optional
|
wenzelm@18695
|
450 |
weight argument for the rule (just as the Pure versions). Weights are
|
wenzelm@18696
|
451 |
ignored by automated tools, but determine the search order of single
|
wenzelm@18694
|
452 |
rule steps.
|
paulson@18557
|
453 |
|
wenzelm@18536
|
454 |
* Syntax: input syntax now supports dummy variable binding "%_. b",
|
wenzelm@18536
|
455 |
where the body does not mention the bound variable. Note that dummy
|
wenzelm@18536
|
456 |
patterns implicitly depend on their context of bounds, which makes
|
wenzelm@18536
|
457 |
"{_. _}" match any set comprehension as expected. Potential
|
wenzelm@18536
|
458 |
INCOMPATIBILITY -- parse translations need to cope with syntactic
|
wenzelm@18536
|
459 |
constant "_idtdummy" in the binding position.
|
wenzelm@18536
|
460 |
|
wenzelm@18536
|
461 |
* Syntax: removed obsolete syntactic constant "_K" and its associated
|
wenzelm@18536
|
462 |
parse translation. INCOMPATIBILITY -- use dummy abstraction instead,
|
wenzelm@18536
|
463 |
for example "A -> B" => "Pi A (%_. B)".
|
wenzelm@17779
|
464 |
|
wenzelm@20582
|
465 |
* Pure: 'class_deps' command visualizes the subclass relation, using
|
wenzelm@20582
|
466 |
the graph browser tool.
|
wenzelm@20582
|
467 |
|
wenzelm@20620
|
468 |
* Pure: 'print_theory' now suppresses entities with internal name
|
wenzelm@20620
|
469 |
(trailing "_") by default; use '!' option for full details.
|
wenzelm@20620
|
470 |
|
wenzelm@17865
|
471 |
|
nipkow@17806
|
472 |
*** HOL ***
|
nipkow@17806
|
473 |
|
haftmann@20607
|
474 |
* New theory OperationalEquality providing class 'eq' with constant 'eq',
|
haftmann@20607
|
475 |
allowing for code generation with polymorphic equality.
|
haftmann@20607
|
476 |
|
haftmann@20485
|
477 |
* Numeral syntax: type 'bin' which was a mere type copy of 'int' has been
|
haftmann@20500
|
478 |
abandoned in favour of plain 'int'. INCOMPATIBILITY -- Significant changes
|
haftmann@20500
|
479 |
for setting up numeral syntax for types:
|
haftmann@20485
|
480 |
|
haftmann@20485
|
481 |
- new constants Numeral.pred and Numeral.succ instead
|
haftmann@20485
|
482 |
of former Numeral.bin_pred and Numeral.bin_succ.
|
haftmann@20485
|
483 |
- Use integer operations instead of bin_add, bin_mult and so on.
|
haftmann@20485
|
484 |
- Numeral simplification theorems named Numeral.numeral_simps instead of Bin_simps.
|
haftmann@20485
|
485 |
- ML structure Bin_Simprocs now named Int_Numeral_Base_Simprocs.
|
haftmann@20485
|
486 |
|
haftmann@20485
|
487 |
See HOL/Integ/IntArith.thy for an example setup.
|
haftmann@20485
|
488 |
|
nipkow@19895
|
489 |
* New top level command 'normal_form' computes the normal form of a term
|
nipkow@19895
|
490 |
that may contain free variables. For example 'normal_form "rev[a,b,c]"'
|
nipkow@19895
|
491 |
prints '[b,c,a]'. This command is suitable for heavy-duty computations
|
nipkow@19895
|
492 |
because the functions are compiled to ML first.
|
nipkow@19895
|
493 |
INCOMPATIBILITY: new keywords 'normal_form' must quoted when used as
|
nipkow@19895
|
494 |
an identifier.
|
nipkow@19895
|
495 |
|
wenzelm@17996
|
496 |
* Alternative iff syntax "A <-> B" for equality on bool (with priority
|
wenzelm@17996
|
497 |
25 like -->); output depends on the "iff" print_mode, the default is
|
wenzelm@17996
|
498 |
"A = B" (with priority 50).
|
wenzelm@17996
|
499 |
|
ballarin@19279
|
500 |
* Renamed constants in HOL.thy and Orderings.thy:
|
haftmann@19233
|
501 |
op + ~> HOL.plus
|
haftmann@19233
|
502 |
op - ~> HOL.minus
|
haftmann@19233
|
503 |
uminus ~> HOL.uminus
|
haftmann@19233
|
504 |
op * ~> HOL.times
|
haftmann@19277
|
505 |
op < ~> Orderings.less
|
haftmann@19277
|
506 |
op <= ~> Orderings.less_eq
|
haftmann@19233
|
507 |
|
haftmann@19233
|
508 |
Adaptions may be required in the following cases:
|
haftmann@19233
|
509 |
|
nipkow@19377
|
510 |
a) User-defined constants using any of the names "plus", "minus", "times",
|
nipkow@19377
|
511 |
"less" or "less_eq". The standard syntax translations for "+", "-" and "*"
|
nipkow@19377
|
512 |
may go wrong.
|
haftmann@19233
|
513 |
INCOMPATIBILITY: use more specific names.
|
haftmann@19233
|
514 |
|
haftmann@19277
|
515 |
b) Variables named "plus", "minus", "times", "less", "less_eq"
|
haftmann@19233
|
516 |
INCOMPATIBILITY: use more specific names.
|
haftmann@19233
|
517 |
|
nipkow@19377
|
518 |
c) Permutative equations (e.g. "a + b = b + a")
|
nipkow@19377
|
519 |
Since the change of names also changes the order of terms, permutative
|
nipkow@19377
|
520 |
rewrite rules may get applied in a different order. Experience shows that
|
nipkow@19377
|
521 |
this is rarely the case (only two adaptions in the whole Isabelle
|
nipkow@19377
|
522 |
distribution).
|
nipkow@19377
|
523 |
INCOMPATIBILITY: rewrite proofs
|
haftmann@19233
|
524 |
|
haftmann@19233
|
525 |
d) ML code directly refering to constant names
|
haftmann@19233
|
526 |
This in general only affects hand-written proof tactics, simprocs and so on.
|
haftmann@19233
|
527 |
INCOMPATIBILITY: grep your sourcecode and replace names.
|
haftmann@19233
|
528 |
|
nipkow@18674
|
529 |
* "LEAST x:A. P" expands to "LEAST x. x:A & P" (input only).
|
nipkow@18674
|
530 |
|
nipkow@19377
|
531 |
* The old set interval syntax "{m..n(}" (and relatives) has been removed.
|
nipkow@19377
|
532 |
Use "{m..<n}" (and relatives) instead.
|
nipkow@19377
|
533 |
|
wenzelm@17865
|
534 |
* In the context of the assumption "~(s = t)" the Simplifier rewrites
|
wenzelm@17865
|
535 |
"t = s" to False (by simproc "neq_simproc"). For backward
|
wenzelm@17865
|
536 |
compatibility this can be disabled by ML "reset use_neq_simproc".
|
wenzelm@17779
|
537 |
|
nipkow@18979
|
538 |
* "m dvd n" where m and n are numbers is evaluated to True/False by simp.
|
nipkow@18979
|
539 |
|
nipkow@19211
|
540 |
* Theorem Cons_eq_map_conv no longer has attribute `simp'.
|
nipkow@19211
|
541 |
|
ballarin@19279
|
542 |
* Theorem setsum_mult renamed to setsum_right_distrib.
|
ballarin@19279
|
543 |
|
nipkow@19211
|
544 |
* Prefer ex1I over ex_ex1I in single-step reasoning, e.g. by the
|
nipkow@19211
|
545 |
'rule' method.
|
nipkow@19211
|
546 |
|
webertj@17809
|
547 |
* Tactics 'sat' and 'satx' reimplemented, several improvements: goals
|
webertj@17809
|
548 |
no longer need to be stated as "<prems> ==> False", equivalences (i.e.
|
wenzelm@17865
|
549 |
"=" on type bool) are handled, variable names of the form "lit_<n>"
|
wenzelm@17865
|
550 |
are no longer reserved, significant speedup.
|
wenzelm@17865
|
551 |
|
webertj@20136
|
552 |
* Tactics 'sat' and 'satx' can now replay MiniSat proof traces. zChaff is
|
webertj@20136
|
553 |
still supported as well.
|
webertj@20136
|
554 |
|
wenzelm@18480
|
555 |
* inductive and datatype: provide projections of mutual rules, bundled
|
wenzelm@18480
|
556 |
as foo_bar.inducts;
|
wenzelm@18480
|
557 |
|
wenzelm@19572
|
558 |
* Library: theory Accessible_Part has been move to main HOL.
|
wenzelm@19572
|
559 |
|
wenzelm@18446
|
560 |
* Library: added theory Coinductive_List of potentially infinite lists
|
wenzelm@18446
|
561 |
as greatest fixed-point.
|
wenzelm@18399
|
562 |
|
wenzelm@19254
|
563 |
* Library: added theory AssocList which implements (finite) maps as
|
schirmer@19252
|
564 |
association lists.
|
webertj@17809
|
565 |
|
berghofe@19855
|
566 |
* New proof method "evaluation" for efficiently solving a goal
|
berghofe@19855
|
567 |
(i.e. a boolean expression) by compiling it to ML. The goal is
|
berghofe@19855
|
568 |
"proved" (via the oracle "Evaluation") if it evaluates to True.
|
berghofe@19855
|
569 |
|
webertj@20217
|
570 |
* Linear arithmetic now splits certain operators (e.g. min, max, abs) also
|
webertj@20217
|
571 |
when invoked by the simplifier. This results in the simplifier being more
|
webertj@20217
|
572 |
powerful on arithmetic goals.
|
webertj@20217
|
573 |
INCOMPATIBILTY: rewrite proofs. Set fast_arith_split_limit to 0 to obtain
|
webertj@20217
|
574 |
the old behavior.
|
webertj@20217
|
575 |
|
kleing@20067
|
576 |
* Support for hex (0x20) and binary (0b1001) numerals.
|
wenzelm@19254
|
577 |
|
chaieb@20327
|
578 |
* New method:
|
chaieb@20327
|
579 |
reify eqs (t), where eqs are equations for an interpretation
|
chaieb@20327
|
580 |
I :: 'a list => 'b => 'c and t::'c is an optional parameter,
|
chaieb@20327
|
581 |
computes a term s::'b and a list xs::'a list and proves the theorem
|
chaieb@20327
|
582 |
I xs s = t. This is also known as reification or quoting. The resulting theorem is applied to the subgoal to substitute t with I xs s.
|
chaieb@20327
|
583 |
If t is omitted, the subgoal itself is reified.
|
chaieb@20327
|
584 |
|
chaieb@20327
|
585 |
* New method:
|
chaieb@20327
|
586 |
reflection corr_thm eqs (t). The parameters eqs and (t) are as explained above. corr_thm is a theorem for
|
chaieb@20327
|
587 |
I vs (f t) = I vs t, where f is supposed to be a computable function (in the sense of code generattion). The method uses reify to compute s and xs as above then applies corr_thm and uses normalization by evaluation to "prove" f s = r and finally gets the theorem t = r, which is again applied to the subgoal. An Example is available in HOL/ex/ReflectionEx.thy.
|
chaieb@20327
|
588 |
|
chaieb@20375
|
589 |
* Reflection: Automatic refification now handels binding, an example is available in HOL/ex/ReflectionEx.thy
|
ballarin@20169
|
590 |
*** HOL-Algebra ***
|
ballarin@20169
|
591 |
|
ballarin@20169
|
592 |
* Method algebra is now set up via an attribute. For examples see CRing.thy.
|
ballarin@20318
|
593 |
INCOMPATIBILITY: the method is now weaker on combinations of algebraic
|
ballarin@20318
|
594 |
structures.
|
ballarin@20318
|
595 |
|
ballarin@20318
|
596 |
* Formalisation of ideals and the quotient construction over rings, contributed
|
ballarin@20318
|
597 |
by Stephan Hohe.
|
ballarin@20318
|
598 |
|
ballarin@20318
|
599 |
* Renamed `CRing.thy' to `Ring.thy'. INCOMPATIBILITY.
|
ballarin@20169
|
600 |
|
wenzelm@19653
|
601 |
*** HOL-Complex ***
|
wenzelm@19653
|
602 |
|
wenzelm@19653
|
603 |
* Theory Real: new method ferrack implements quantifier elimination
|
wenzelm@19653
|
604 |
for linear arithmetic over the reals. The quantifier elimination
|
wenzelm@19653
|
605 |
feature is used only for decision, for compatibility with arith. This
|
wenzelm@19653
|
606 |
means a goal is either solved or left unchanged, no simplification.
|
wenzelm@19653
|
607 |
|
wenzelm@19653
|
608 |
|
wenzelm@17878
|
609 |
*** ML ***
|
wenzelm@17878
|
610 |
|
haftmann@20607
|
611 |
* Pure/General/susp.ML:
|
haftmann@20607
|
612 |
|
haftmann@20607
|
613 |
New module for delayed evaluations.
|
haftmann@20607
|
614 |
|
haftmann@18450
|
615 |
* Pure/library:
|
haftmann@18450
|
616 |
|
haftmann@20348
|
617 |
Semantically identical functions "equal_list" and "eq_list" have been
|
haftmann@20348
|
618 |
unified to "eq_list".
|
haftmann@20348
|
619 |
|
haftmann@20348
|
620 |
* Pure/library:
|
haftmann@20348
|
621 |
|
haftmann@18450
|
622 |
val burrow: ('a list -> 'b list) -> 'a list list -> 'b list list
|
haftmann@18549
|
623 |
val fold_burrow: ('a list -> 'c -> 'b list * 'd) -> 'a list list -> 'c -> 'b list list * 'd
|
haftmann@18450
|
624 |
|
wenzelm@18540
|
625 |
The semantics of "burrow" is: "take a function with *simulatanously*
|
wenzelm@18540
|
626 |
transforms a list of value, and apply it *simulatanously* to a list of
|
wenzelm@18540
|
627 |
list of values of the appropriate type". Confer this with "map" which
|
wenzelm@18540
|
628 |
would *not* apply its argument function simulatanously but in
|
haftmann@18549
|
629 |
sequence. "fold_burrow" has an additional context.
|
wenzelm@18540
|
630 |
|
wenzelm@18540
|
631 |
Both actually avoid the usage of "unflat" since they hide away
|
wenzelm@18540
|
632 |
"unflat" from the user.
|
haftmann@18450
|
633 |
|
wenzelm@18446
|
634 |
* Pure/library: functions map2 and fold2 with curried syntax for
|
wenzelm@18446
|
635 |
simultanous mapping and folding:
|
wenzelm@18446
|
636 |
|
haftmann@18422
|
637 |
val map2: ('a -> 'b -> 'c) -> 'a list -> 'b list -> 'c list
|
haftmann@18422
|
638 |
val fold2: ('a -> 'b -> 'c -> 'c) -> 'a list -> 'b list -> 'c -> 'c
|
haftmann@18422
|
639 |
|
wenzelm@18446
|
640 |
* Pure/library: indexed lists - some functions in the Isabelle library
|
wenzelm@18446
|
641 |
treating lists over 'a as finite mappings from [0...n] to 'a have been
|
wenzelm@18446
|
642 |
given more convenient names and signatures reminiscent of similar
|
wenzelm@18446
|
643 |
functions for alists, tables, etc:
|
haftmann@18051
|
644 |
|
haftmann@18051
|
645 |
val nth: 'a list -> int -> 'a
|
haftmann@18051
|
646 |
val nth_update: int * 'a -> 'a list -> 'a list
|
haftmann@18051
|
647 |
val nth_map: int -> ('a -> 'a) -> 'a list -> 'a list
|
haftmann@18051
|
648 |
val fold_index: (int * 'a -> 'b -> 'b) -> 'a list -> 'b -> 'b
|
haftmann@18051
|
649 |
|
wenzelm@18446
|
650 |
Note that fold_index starts counting at index 0, not 1 like foldln
|
wenzelm@18446
|
651 |
used to.
|
wenzelm@18446
|
652 |
|
wenzelm@19653
|
653 |
* Pure/library: general ``divide_and_conquer'' combinator on lists.
|
wenzelm@19653
|
654 |
|
wenzelm@19032
|
655 |
* Pure/General/name_mangler.ML provides a functor for generic name
|
wenzelm@18446
|
656 |
mangling (bijective mapping from any expression values to strings).
|
wenzelm@18446
|
657 |
|
wenzelm@19032
|
658 |
* Pure/General/rat.ML implements rational numbers.
|
wenzelm@19032
|
659 |
|
wenzelm@19032
|
660 |
* Pure/General/table.ML: the join operations now works via exceptions
|
wenzelm@19081
|
661 |
DUP/SAME instead of type option. This is simpler in simple cases, and
|
wenzelm@19081
|
662 |
admits slightly more efficient complex applications.
|
wenzelm@18446
|
663 |
|
wenzelm@18642
|
664 |
* Pure: datatype Context.generic joins theory/Proof.context and
|
wenzelm@18644
|
665 |
provides some facilities for code that works in either kind of
|
wenzelm@18642
|
666 |
context, notably GenericDataFun for uniform theory and proof data.
|
wenzelm@18642
|
667 |
|
wenzelm@18862
|
668 |
* Pure: 'advanced' translation functions (parse_translation etc.) now
|
wenzelm@18862
|
669 |
use Context.generic instead of just theory.
|
wenzelm@18862
|
670 |
|
wenzelm@18737
|
671 |
* Pure: simplified internal attribute type, which is now always
|
wenzelm@18737
|
672 |
Context.generic * thm -> Context.generic * thm. Global (theory)
|
wenzelm@18737
|
673 |
vs. local (Proof.context) attributes have been discontinued, while
|
wenzelm@18738
|
674 |
minimizing code duplication. Thm.rule_attribute and
|
wenzelm@18738
|
675 |
Thm.declaration_attribute build canonical attributes; see also
|
wenzelm@19006
|
676 |
structure Context for further operations on Context.generic, notably
|
wenzelm@19006
|
677 |
GenericDataFun. INCOMPATIBILITY, need to adapt attribute type
|
wenzelm@19006
|
678 |
declarations and definitions.
|
wenzelm@19006
|
679 |
|
wenzelm@19508
|
680 |
* Pure/kernel: consts certification ignores sort constraints given in
|
wenzelm@19508
|
681 |
signature declarations. (This information is not relevant to the
|
wenzelm@19508
|
682 |
logic, but only for type inference.) IMPORTANT INTERNAL CHANGE.
|
wenzelm@19508
|
683 |
|
wenzelm@19508
|
684 |
* Pure: axiomatic type classes are now purely definitional, with
|
wenzelm@19508
|
685 |
explicit proofs of class axioms and super class relations performed
|
wenzelm@19508
|
686 |
internally. See Pure/axclass.ML for the main internal interfaces --
|
wenzelm@19508
|
687 |
notably AxClass.define_class supercedes AxClass.add_axclass, and
|
wenzelm@19508
|
688 |
AxClass.axiomatize_class/classrel/arity supercede
|
wenzelm@19508
|
689 |
Sign.add_classes/classrel/arities.
|
wenzelm@19508
|
690 |
|
wenzelm@19006
|
691 |
* Pure/Isar: Args/Attrib parsers operate on Context.generic --
|
wenzelm@19006
|
692 |
global/local versions on theory vs. Proof.context have been
|
wenzelm@19006
|
693 |
discontinued; Attrib.syntax and Method.syntax have been adapted
|
wenzelm@19006
|
694 |
accordingly. INCOMPATIBILITY, need to adapt parser expressions for
|
wenzelm@19006
|
695 |
attributes, methods, etc.
|
wenzelm@18642
|
696 |
|
wenzelm@18446
|
697 |
* Pure: several functions of signature "... -> theory -> theory * ..."
|
wenzelm@18446
|
698 |
have been reoriented to "... -> theory -> ... * theory" in order to
|
wenzelm@18446
|
699 |
allow natural usage in combination with the ||>, ||>>, |-> and
|
wenzelm@18446
|
700 |
fold_map combinators.
|
haftmann@18051
|
701 |
|
wenzelm@18020
|
702 |
* Pure: primitive rule lift_rule now takes goal cterm instead of an
|
wenzelm@18145
|
703 |
actual goal state (thm). Use Thm.lift_rule (Thm.cprem_of st i) to
|
wenzelm@18020
|
704 |
achieve the old behaviour.
|
wenzelm@18020
|
705 |
|
wenzelm@18020
|
706 |
* Pure: the "Goal" constant is now called "prop", supporting a
|
wenzelm@18020
|
707 |
slightly more general idea of ``protecting'' meta-level rule
|
wenzelm@18020
|
708 |
statements.
|
wenzelm@18020
|
709 |
|
wenzelm@20040
|
710 |
* Pure: Logic.(un)varify only works in a global context, which is now
|
wenzelm@20040
|
711 |
enforced instead of silently assumed. INCOMPATIBILITY, may use
|
wenzelm@20040
|
712 |
Logic.legacy_(un)varify as temporary workaround.
|
wenzelm@20040
|
713 |
|
wenzelm@20090
|
714 |
* Pure: structure Name provides scalable operations for generating
|
wenzelm@20090
|
715 |
internal variable names, notably Name.variants etc. This replaces
|
wenzelm@20090
|
716 |
some popular functions from term.ML:
|
wenzelm@20090
|
717 |
|
wenzelm@20090
|
718 |
Term.variant -> Name.variant
|
wenzelm@20090
|
719 |
Term.variantlist -> Name.variant_list (*canonical argument order*)
|
wenzelm@20090
|
720 |
Term.invent_names -> Name.invent_list
|
wenzelm@20090
|
721 |
|
wenzelm@20090
|
722 |
Note that low-level renaming rarely occurs in new code -- operations
|
wenzelm@20090
|
723 |
from structure Variable are used instead (see below).
|
wenzelm@20090
|
724 |
|
wenzelm@20040
|
725 |
* Pure: structure Variable provides fundamental operations for proper
|
wenzelm@20040
|
726 |
treatment of fixed/schematic variables in a context. For example,
|
wenzelm@20040
|
727 |
Variable.import introduces fixes for schematics of given facts and
|
wenzelm@20040
|
728 |
Variable.export reverses the effect (up to renaming) -- this replaces
|
wenzelm@20040
|
729 |
various freeze_thaw operations.
|
wenzelm@20040
|
730 |
|
wenzelm@18567
|
731 |
* Pure: structure Goal provides simple interfaces for
|
wenzelm@17981
|
732 |
init/conclude/finish and tactical prove operations (replacing former
|
wenzelm@20040
|
733 |
Tactic.prove). Goal.prove is the canonical way to prove results
|
wenzelm@20040
|
734 |
within a given context; Goal.prove_global is a degraded version for
|
wenzelm@20040
|
735 |
theory level goals, including a global Drule.standard. Note that
|
wenzelm@20040
|
736 |
OldGoals.prove_goalw_cterm has long been obsolete, since it is
|
wenzelm@20040
|
737 |
ill-behaved in a local proof context (e.g. with local fixes/assumes or
|
wenzelm@20040
|
738 |
in a locale context).
|
wenzelm@17981
|
739 |
|
wenzelm@18815
|
740 |
* Isar: simplified treatment of user-level errors, using exception
|
wenzelm@18687
|
741 |
ERROR of string uniformly. Function error now merely raises ERROR,
|
wenzelm@18686
|
742 |
without any side effect on output channels. The Isar toplevel takes
|
wenzelm@18686
|
743 |
care of proper display of ERROR exceptions. ML code may use plain
|
wenzelm@18686
|
744 |
handle/can/try; cat_error may be used to concatenate errors like this:
|
wenzelm@18686
|
745 |
|
wenzelm@18686
|
746 |
... handle ERROR msg => cat_error msg "..."
|
wenzelm@18686
|
747 |
|
wenzelm@18686
|
748 |
Toplevel ML code (run directly or through the Isar toplevel) may be
|
wenzelm@18687
|
749 |
embedded into the Isar toplevel with exception display/debug like
|
wenzelm@18687
|
750 |
this:
|
wenzelm@18686
|
751 |
|
wenzelm@18686
|
752 |
Isar.toplevel (fn () => ...)
|
wenzelm@18686
|
753 |
|
wenzelm@18686
|
754 |
INCOMPATIBILITY, removed special transform_error facilities, removed
|
wenzelm@18686
|
755 |
obsolete variants of user-level exceptions (ERROR_MESSAGE,
|
wenzelm@18686
|
756 |
Context.PROOF, ProofContext.CONTEXT, Proof.STATE, ProofHistory.FAIL)
|
wenzelm@18686
|
757 |
-- use plain ERROR instead.
|
wenzelm@18686
|
758 |
|
wenzelm@18815
|
759 |
* Isar: theory setup now has type (theory -> theory), instead of a
|
wenzelm@18722
|
760 |
list. INCOMPATIBILITY, may use #> to compose setup functions.
|
wenzelm@18722
|
761 |
|
wenzelm@18815
|
762 |
* Isar: installed ML toplevel pretty printer for type Proof.context,
|
wenzelm@18815
|
763 |
subject to ProofContext.debug/verbose flags.
|
wenzelm@18815
|
764 |
|
wenzelm@18815
|
765 |
* Isar: Toplevel.theory_to_proof admits transactions that modify the
|
wenzelm@18815
|
766 |
theory before entering a proof state. Transactions now always see a
|
wenzelm@18815
|
767 |
quasi-functional intermediate checkpoint, both in interactive and
|
wenzelm@18590
|
768 |
batch mode.
|
wenzelm@18567
|
769 |
|
wenzelm@17878
|
770 |
* Simplifier: the simpset of a running simplification process now
|
wenzelm@17878
|
771 |
contains a proof context (cf. Simplifier.the_context), which is the
|
wenzelm@17878
|
772 |
very context that the initial simpset has been retrieved from (by
|
wenzelm@17890
|
773 |
simpset_of/local_simpset_of). Consequently, all plug-in components
|
wenzelm@17878
|
774 |
(solver, looper etc.) may depend on arbitrary proof data.
|
wenzelm@17878
|
775 |
|
wenzelm@17878
|
776 |
* Simplifier.inherit_context inherits the proof context (plus the
|
wenzelm@17878
|
777 |
local bounds) of the current simplification process; any simproc
|
wenzelm@17878
|
778 |
etc. that calls the Simplifier recursively should do this! Removed
|
wenzelm@17878
|
779 |
former Simplifier.inherit_bounds, which is already included here --
|
wenzelm@17890
|
780 |
INCOMPATIBILITY. Tools based on low-level rewriting may even have to
|
wenzelm@17890
|
781 |
specify an explicit context using Simplifier.context/theory_context.
|
wenzelm@17878
|
782 |
|
wenzelm@17878
|
783 |
* Simplifier/Classical Reasoner: more abstract interfaces
|
wenzelm@17878
|
784 |
change_simpset/claset for modifying the simpset/claset reference of a
|
wenzelm@17878
|
785 |
theory; raw versions simpset/claset_ref etc. have been discontinued --
|
wenzelm@17878
|
786 |
INCOMPATIBILITY.
|
wenzelm@17878
|
787 |
|
wenzelm@18540
|
788 |
* Provers: more generic wrt. syntax of object-logics, avoid hardwired
|
wenzelm@18540
|
789 |
"Trueprop" etc.
|
wenzelm@18540
|
790 |
|
wenzelm@17878
|
791 |
|
wenzelm@17754
|
792 |
|
wenzelm@17720
|
793 |
New in Isabelle2005 (October 2005)
|
wenzelm@17720
|
794 |
----------------------------------
|
wenzelm@14655
|
795 |
|
wenzelm@14655
|
796 |
*** General ***
|
wenzelm@14655
|
797 |
|
nipkow@15130
|
798 |
* Theory headers: the new header syntax for Isar theories is
|
nipkow@15130
|
799 |
|
nipkow@15130
|
800 |
theory <name>
|
wenzelm@16234
|
801 |
imports <theory1> ... <theoryN>
|
wenzelm@16234
|
802 |
uses <file1> ... <fileM>
|
nipkow@15130
|
803 |
begin
|
nipkow@15130
|
804 |
|
wenzelm@16234
|
805 |
where the 'uses' part is optional. The previous syntax
|
wenzelm@16234
|
806 |
|
wenzelm@16234
|
807 |
theory <name> = <theory1> + ... + <theoryN>:
|
wenzelm@16234
|
808 |
|
wenzelm@16717
|
809 |
will disappear in the next release. Use isatool fixheaders to convert
|
wenzelm@16717
|
810 |
existing theory files. Note that there is no change in ancient
|
wenzelm@17371
|
811 |
non-Isar theories now, but these will disappear soon.
|
nipkow@15130
|
812 |
|
berghofe@15475
|
813 |
* Theory loader: parent theories can now also be referred to via
|
wenzelm@16234
|
814 |
relative and absolute paths.
|
wenzelm@16234
|
815 |
|
wenzelm@17408
|
816 |
* Command 'find_theorems' searches for a list of criteria instead of a
|
wenzelm@17408
|
817 |
list of constants. Known criteria are: intro, elim, dest, name:string,
|
wenzelm@17408
|
818 |
simp:term, and any term. Criteria can be preceded by '-' to select
|
wenzelm@17408
|
819 |
theorems that do not match. Intro, elim, dest select theorems that
|
wenzelm@17408
|
820 |
match the current goal, name:s selects theorems whose fully qualified
|
wenzelm@17408
|
821 |
name contain s, and simp:term selects all simplification rules whose
|
wenzelm@17408
|
822 |
lhs match term. Any other term is interpreted as pattern and selects
|
wenzelm@17408
|
823 |
all theorems matching the pattern. Available in ProofGeneral under
|
wenzelm@17408
|
824 |
'ProofGeneral -> Find Theorems' or C-c C-f. Example:
|
wenzelm@16234
|
825 |
|
wenzelm@17275
|
826 |
C-c C-f (100) "(_::nat) + _ + _" intro -name: "HOL."
|
wenzelm@16234
|
827 |
|
wenzelm@16234
|
828 |
prints the last 100 theorems matching the pattern "(_::nat) + _ + _",
|
wenzelm@16234
|
829 |
matching the current goal as introduction rule and not having "HOL."
|
wenzelm@16234
|
830 |
in their name (i.e. not being defined in theory HOL).
|
wenzelm@16013
|
831 |
|
wenzelm@17408
|
832 |
* Command 'thms_containing' has been discontinued in favour of
|
wenzelm@17408
|
833 |
'find_theorems'; INCOMPATIBILITY.
|
wenzelm@17408
|
834 |
|
wenzelm@17385
|
835 |
* Communication with Proof General is now 8bit clean, which means that
|
wenzelm@17385
|
836 |
Unicode text in UTF-8 encoding may be used within theory texts (both
|
wenzelm@17408
|
837 |
formal and informal parts). Cf. option -U of the Isabelle Proof
|
wenzelm@17538
|
838 |
General interface. Here are some simple examples (cf. src/HOL/ex):
|
wenzelm@17538
|
839 |
|
wenzelm@17538
|
840 |
http://isabelle.in.tum.de/library/HOL/ex/Hebrew.html
|
wenzelm@17538
|
841 |
http://isabelle.in.tum.de/library/HOL/ex/Chinese.html
|
wenzelm@17385
|
842 |
|
wenzelm@17425
|
843 |
* Improved efficiency of the Simplifier and, to a lesser degree, the
|
wenzelm@17425
|
844 |
Classical Reasoner. Typical big applications run around 2 times
|
wenzelm@17425
|
845 |
faster.
|
wenzelm@17425
|
846 |
|
wenzelm@15703
|
847 |
|
wenzelm@15703
|
848 |
*** Document preparation ***
|
wenzelm@15703
|
849 |
|
wenzelm@16234
|
850 |
* Commands 'display_drafts' and 'print_drafts' perform simple output
|
wenzelm@16234
|
851 |
of raw sources. Only those symbols that do not require additional
|
wenzelm@16234
|
852 |
LaTeX packages (depending on comments in isabellesym.sty) are
|
wenzelm@16234
|
853 |
displayed properly, everything else is left verbatim. isatool display
|
wenzelm@16234
|
854 |
and isatool print are used as front ends (these are subject to the
|
wenzelm@16234
|
855 |
DVI/PDF_VIEWER and PRINT_COMMAND settings, respectively).
|
wenzelm@16234
|
856 |
|
wenzelm@17047
|
857 |
* Command tags control specific markup of certain regions of text,
|
wenzelm@17047
|
858 |
notably folding and hiding. Predefined tags include "theory" (for
|
wenzelm@17047
|
859 |
theory begin and end), "proof" for proof commands, and "ML" for
|
wenzelm@17047
|
860 |
commands involving ML code; the additional tags "visible" and
|
wenzelm@17047
|
861 |
"invisible" are unused by default. Users may give explicit tag
|
wenzelm@17047
|
862 |
specifications in the text, e.g. ''by %invisible (auto)''. The
|
wenzelm@17047
|
863 |
interpretation of tags is determined by the LaTeX job during document
|
wenzelm@17047
|
864 |
preparation: see option -V of isatool usedir, or options -n and -t of
|
wenzelm@17047
|
865 |
isatool document, or even the LaTeX macros \isakeeptag, \isafoldtag,
|
wenzelm@17047
|
866 |
\isadroptag.
|
wenzelm@17047
|
867 |
|
wenzelm@17047
|
868 |
Several document versions may be produced at the same time via isatool
|
wenzelm@17047
|
869 |
usedir (the generated index.html will link all of them). Typical
|
wenzelm@17047
|
870 |
specifications include ''-V document=theory,proof,ML'' to present
|
wenzelm@17047
|
871 |
theory/proof/ML parts faithfully, ''-V outline=/proof,/ML'' to fold
|
wenzelm@17047
|
872 |
proof and ML commands, and ''-V mutilated=-theory,-proof,-ML'' to omit
|
wenzelm@17047
|
873 |
these parts without any formal replacement text. The Isabelle site
|
wenzelm@17047
|
874 |
default settings produce ''document'' and ''outline'' versions as
|
wenzelm@17047
|
875 |
specified above.
|
wenzelm@16234
|
876 |
|
haftmann@17402
|
877 |
* Several new antiquotations:
|
wenzelm@15979
|
878 |
|
wenzelm@15979
|
879 |
@{term_type term} prints a term with its type annotated;
|
wenzelm@15979
|
880 |
|
wenzelm@15979
|
881 |
@{typeof term} prints the type of a term;
|
wenzelm@15979
|
882 |
|
wenzelm@16234
|
883 |
@{const const} is the same as @{term const}, but checks that the
|
wenzelm@16234
|
884 |
argument is a known logical constant;
|
wenzelm@15979
|
885 |
|
wenzelm@15979
|
886 |
@{term_style style term} and @{thm_style style thm} print a term or
|
wenzelm@16234
|
887 |
theorem applying a "style" to it
|
wenzelm@16234
|
888 |
|
wenzelm@17117
|
889 |
@{ML text}
|
wenzelm@17117
|
890 |
|
wenzelm@16234
|
891 |
Predefined styles are 'lhs' and 'rhs' printing the lhs/rhs of
|
wenzelm@16234
|
892 |
definitions, equations, inequations etc., 'concl' printing only the
|
schirmer@17393
|
893 |
conclusion of a meta-logical statement theorem, and 'prem1' .. 'prem19'
|
wenzelm@16234
|
894 |
to print the specified premise. TermStyle.add_style provides an ML
|
wenzelm@16234
|
895 |
interface for introducing further styles. See also the "LaTeX Sugar"
|
wenzelm@17117
|
896 |
document practical applications. The ML antiquotation prints
|
wenzelm@17117
|
897 |
type-checked ML expressions verbatim.
|
wenzelm@16234
|
898 |
|
wenzelm@17259
|
899 |
* Markup commands 'chapter', 'section', 'subsection', 'subsubsection',
|
wenzelm@17259
|
900 |
and 'text' support optional locale specification '(in loc)', which
|
wenzelm@17269
|
901 |
specifies the default context for interpreting antiquotations. For
|
wenzelm@17269
|
902 |
example: 'text (in lattice) {* @{thm inf_assoc}*}'.
|
wenzelm@17259
|
903 |
|
wenzelm@17259
|
904 |
* Option 'locale=NAME' of antiquotations specifies an alternative
|
wenzelm@17259
|
905 |
context interpreting the subsequent argument. For example: @{thm
|
wenzelm@17269
|
906 |
[locale=lattice] inf_assoc}.
|
wenzelm@17259
|
907 |
|
wenzelm@17097
|
908 |
* Proper output of proof terms (@{prf ...} and @{full_prf ...}) within
|
wenzelm@17097
|
909 |
a proof context.
|
wenzelm@17097
|
910 |
|
wenzelm@17097
|
911 |
* Proper output of antiquotations for theory commands involving a
|
wenzelm@17097
|
912 |
proof context (such as 'locale' or 'theorem (in loc) ...').
|
wenzelm@17097
|
913 |
|
wenzelm@17193
|
914 |
* Delimiters of outer tokens (string etc.) now produce separate LaTeX
|
wenzelm@17193
|
915 |
macros (\isachardoublequoteopen, isachardoublequoteclose etc.).
|
wenzelm@17193
|
916 |
|
wenzelm@17193
|
917 |
* isatool usedir: new option -C (default true) controls whether option
|
wenzelm@17193
|
918 |
-D should include a copy of the original document directory; -C false
|
wenzelm@17193
|
919 |
prevents unwanted effects such as copying of administrative CVS data.
|
wenzelm@17193
|
920 |
|
wenzelm@16234
|
921 |
|
wenzelm@16234
|
922 |
*** Pure ***
|
wenzelm@16234
|
923 |
|
wenzelm@16234
|
924 |
* Considerably improved version of 'constdefs' command. Now performs
|
wenzelm@16234
|
925 |
automatic type-inference of declared constants; additional support for
|
wenzelm@16234
|
926 |
local structure declarations (cf. locales and HOL records), see also
|
wenzelm@16234
|
927 |
isar-ref manual. Potential INCOMPATIBILITY: need to observe strictly
|
wenzelm@16234
|
928 |
sequential dependencies of definitions within a single 'constdefs'
|
wenzelm@16234
|
929 |
section; moreover, the declared name needs to be an identifier. If
|
wenzelm@16234
|
930 |
all fails, consider to fall back on 'consts' and 'defs' separately.
|
wenzelm@16234
|
931 |
|
wenzelm@16234
|
932 |
* Improved indexed syntax and implicit structures. First of all,
|
wenzelm@16234
|
933 |
indexed syntax provides a notational device for subscripted
|
wenzelm@16234
|
934 |
application, using the new syntax \<^bsub>term\<^esub> for arbitrary
|
wenzelm@16234
|
935 |
expressions. Secondly, in a local context with structure
|
wenzelm@16234
|
936 |
declarations, number indexes \<^sub>n or the empty index (default
|
wenzelm@16234
|
937 |
number 1) refer to a certain fixed variable implicitly; option
|
wenzelm@16234
|
938 |
show_structs controls printing of implicit structures. Typical
|
wenzelm@16234
|
939 |
applications of these concepts involve record types and locales.
|
wenzelm@16234
|
940 |
|
wenzelm@16234
|
941 |
* New command 'no_syntax' removes grammar declarations (and
|
wenzelm@16234
|
942 |
translations) resulting from the given syntax specification, which is
|
wenzelm@16234
|
943 |
interpreted in the same manner as for the 'syntax' command.
|
wenzelm@16234
|
944 |
|
wenzelm@16234
|
945 |
* 'Advanced' translation functions (parse_translation etc.) may depend
|
wenzelm@16234
|
946 |
on the signature of the theory context being presently used for
|
wenzelm@16234
|
947 |
parsing/printing, see also isar-ref manual.
|
wenzelm@16234
|
948 |
|
wenzelm@16856
|
949 |
* Improved 'oracle' command provides a type-safe interface to turn an
|
wenzelm@16856
|
950 |
ML expression of type theory -> T -> term into a primitive rule of
|
wenzelm@16856
|
951 |
type theory -> T -> thm (i.e. the functionality of Thm.invoke_oracle
|
wenzelm@16856
|
952 |
is already included here); see also FOL/ex/IffExample.thy;
|
wenzelm@16856
|
953 |
INCOMPATIBILITY.
|
wenzelm@16856
|
954 |
|
wenzelm@17275
|
955 |
* axclass: name space prefix for class "c" is now "c_class" (was "c"
|
wenzelm@17275
|
956 |
before); "cI" is no longer bound, use "c.intro" instead.
|
wenzelm@17275
|
957 |
INCOMPATIBILITY. This change avoids clashes of fact bindings for
|
wenzelm@17275
|
958 |
axclasses vs. locales.
|
wenzelm@17275
|
959 |
|
wenzelm@16234
|
960 |
* Improved internal renaming of symbolic identifiers -- attach primes
|
wenzelm@16234
|
961 |
instead of base 26 numbers.
|
wenzelm@16234
|
962 |
|
wenzelm@16234
|
963 |
* New flag show_question_marks controls printing of leading question
|
wenzelm@16234
|
964 |
marks in schematic variable names.
|
wenzelm@16234
|
965 |
|
wenzelm@16234
|
966 |
* In schematic variable names, *any* symbol following \<^isub> or
|
wenzelm@16234
|
967 |
\<^isup> is now treated as part of the base name. For example, the
|
wenzelm@16234
|
968 |
following works without printing of awkward ".0" indexes:
|
wenzelm@16234
|
969 |
|
wenzelm@16234
|
970 |
lemma "x\<^isub>1 = x\<^isub>2 ==> x\<^isub>2 = x\<^isub>1"
|
wenzelm@16234
|
971 |
by simp
|
wenzelm@16234
|
972 |
|
wenzelm@16234
|
973 |
* Inner syntax includes (*(*nested*) comments*).
|
wenzelm@16234
|
974 |
|
wenzelm@17548
|
975 |
* Pretty printer now supports unbreakable blocks, specified in mixfix
|
wenzelm@16234
|
976 |
annotations as "(00...)".
|
wenzelm@16234
|
977 |
|
wenzelm@16234
|
978 |
* Clear separation of logical types and nonterminals, where the latter
|
wenzelm@16234
|
979 |
may only occur in 'syntax' specifications or type abbreviations.
|
wenzelm@16234
|
980 |
Before that distinction was only partially implemented via type class
|
wenzelm@16234
|
981 |
"logic" vs. "{}". Potential INCOMPATIBILITY in rare cases of improper
|
wenzelm@16234
|
982 |
use of 'types'/'consts' instead of 'nonterminals'/'syntax'. Some very
|
wenzelm@16234
|
983 |
exotic syntax specifications may require further adaption
|
wenzelm@17691
|
984 |
(e.g. Cube/Cube.thy).
|
wenzelm@16234
|
985 |
|
wenzelm@16234
|
986 |
* Removed obsolete type class "logic", use the top sort {} instead.
|
wenzelm@16234
|
987 |
Note that non-logical types should be declared as 'nonterminals'
|
wenzelm@16234
|
988 |
rather than 'types'. INCOMPATIBILITY for new object-logic
|
wenzelm@16234
|
989 |
specifications.
|
wenzelm@16234
|
990 |
|
ballarin@17095
|
991 |
* Attributes 'induct' and 'cases': type or set names may now be
|
ballarin@17095
|
992 |
locally fixed variables as well.
|
ballarin@17095
|
993 |
|
wenzelm@16234
|
994 |
* Simplifier: can now control the depth to which conditional rewriting
|
wenzelm@16234
|
995 |
is traced via the PG menu Isabelle -> Settings -> Trace Simp Depth
|
wenzelm@16234
|
996 |
Limit.
|
wenzelm@16234
|
997 |
|
wenzelm@16234
|
998 |
* Simplifier: simplification procedures may now take the current
|
wenzelm@16234
|
999 |
simpset into account (cf. Simplifier.simproc(_i) / mk_simproc
|
wenzelm@16234
|
1000 |
interface), which is very useful for calling the Simplifier
|
wenzelm@16234
|
1001 |
recursively. Minor INCOMPATIBILITY: the 'prems' argument of simprocs
|
wenzelm@16234
|
1002 |
is gone -- use prems_of_ss on the simpset instead. Moreover, the
|
wenzelm@16234
|
1003 |
low-level mk_simproc no longer applies Logic.varify internally, to
|
wenzelm@16234
|
1004 |
allow for use in a context of fixed variables.
|
wenzelm@16234
|
1005 |
|
wenzelm@16234
|
1006 |
* thin_tac now works even if the assumption being deleted contains !!
|
wenzelm@16234
|
1007 |
or ==>. More generally, erule now works even if the major premise of
|
wenzelm@16234
|
1008 |
the elimination rule contains !! or ==>.
|
wenzelm@16234
|
1009 |
|
wenzelm@17597
|
1010 |
* Method 'rules' has been renamed to 'iprover'. INCOMPATIBILITY.
|
nipkow@17590
|
1011 |
|
wenzelm@16234
|
1012 |
* Reorganized bootstrapping of the Pure theories; CPure is now derived
|
wenzelm@16234
|
1013 |
from Pure, which contains all common declarations already. Both
|
wenzelm@16234
|
1014 |
theories are defined via plain Isabelle/Isar .thy files.
|
wenzelm@16234
|
1015 |
INCOMPATIBILITY: elements of CPure (such as the CPure.intro /
|
wenzelm@16234
|
1016 |
CPure.elim / CPure.dest attributes) now appear in the Pure name space;
|
wenzelm@16234
|
1017 |
use isatool fixcpure to adapt your theory and ML sources.
|
wenzelm@16234
|
1018 |
|
wenzelm@16234
|
1019 |
* New syntax 'name(i-j, i-, i, ...)' for referring to specific
|
wenzelm@16234
|
1020 |
selections of theorems in named facts via index ranges.
|
wenzelm@16234
|
1021 |
|
wenzelm@17097
|
1022 |
* 'print_theorems': in theory mode, really print the difference
|
wenzelm@17097
|
1023 |
wrt. the last state (works for interactive theory development only),
|
wenzelm@17097
|
1024 |
in proof mode print all local facts (cf. 'print_facts');
|
wenzelm@17097
|
1025 |
|
wenzelm@17397
|
1026 |
* 'hide': option '(open)' hides only base names.
|
wenzelm@17397
|
1027 |
|
wenzelm@17275
|
1028 |
* More efficient treatment of intermediate checkpoints in interactive
|
wenzelm@17275
|
1029 |
theory development.
|
wenzelm@17275
|
1030 |
|
berghofe@17663
|
1031 |
* Code generator is now invoked via code_module (incremental code
|
wenzelm@17664
|
1032 |
generation) and code_library (modular code generation, ML structures
|
wenzelm@17664
|
1033 |
for each theory). INCOMPATIBILITY: new keywords 'file' and 'contains'
|
wenzelm@17664
|
1034 |
must be quoted when used as identifiers.
|
wenzelm@17664
|
1035 |
|
wenzelm@17664
|
1036 |
* New 'value' command for reading, evaluating and printing terms using
|
wenzelm@17664
|
1037 |
the code generator. INCOMPATIBILITY: command keyword 'value' must be
|
wenzelm@17664
|
1038 |
quoted when used as identifier.
|
berghofe@17663
|
1039 |
|
wenzelm@16234
|
1040 |
|
wenzelm@16234
|
1041 |
*** Locales ***
|
ballarin@17095
|
1042 |
|
wenzelm@17385
|
1043 |
* New commands for the interpretation of locale expressions in
|
wenzelm@17385
|
1044 |
theories (1), locales (2) and proof contexts (3). These generate
|
wenzelm@17385
|
1045 |
proof obligations from the expression specification. After the
|
wenzelm@17385
|
1046 |
obligations have been discharged, theorems of the expression are added
|
wenzelm@17385
|
1047 |
to the theory, target locale or proof context. The synopsis of the
|
wenzelm@17385
|
1048 |
commands is a follows:
|
wenzelm@17385
|
1049 |
|
ballarin@17095
|
1050 |
(1) interpretation expr inst
|
ballarin@17095
|
1051 |
(2) interpretation target < expr
|
ballarin@17095
|
1052 |
(3) interpret expr inst
|
wenzelm@17385
|
1053 |
|
ballarin@17095
|
1054 |
Interpretation in theories and proof contexts require a parameter
|
ballarin@17095
|
1055 |
instantiation of terms from the current context. This is applied to
|
wenzelm@17385
|
1056 |
specifications and theorems of the interpreted expression.
|
wenzelm@17385
|
1057 |
Interpretation in locales only permits parameter renaming through the
|
wenzelm@17385
|
1058 |
locale expression. Interpretation is smart in that interpretations
|
wenzelm@17385
|
1059 |
that are active already do not occur in proof obligations, neither are
|
wenzelm@17385
|
1060 |
instantiated theorems stored in duplicate. Use 'print_interps' to
|
wenzelm@17385
|
1061 |
inspect active interpretations of a particular locale. For details,
|
ballarin@17436
|
1062 |
see the Isar Reference manual. Examples can be found in
|
ballarin@17436
|
1063 |
HOL/Finite_Set.thy and HOL/Algebra/UnivPoly.thy.
|
wenzelm@16234
|
1064 |
|
wenzelm@16234
|
1065 |
INCOMPATIBILITY: former 'instantiate' has been withdrawn, use
|
wenzelm@16234
|
1066 |
'interpret' instead.
|
wenzelm@16234
|
1067 |
|
wenzelm@17385
|
1068 |
* New context element 'constrains' for adding type constraints to
|
wenzelm@17385
|
1069 |
parameters.
|
wenzelm@17385
|
1070 |
|
wenzelm@17385
|
1071 |
* Context expressions: renaming of parameters with syntax
|
wenzelm@17385
|
1072 |
redeclaration.
|
ballarin@17095
|
1073 |
|
ballarin@17095
|
1074 |
* Locale declaration: 'includes' disallowed.
|
ballarin@17095
|
1075 |
|
wenzelm@16234
|
1076 |
* Proper static binding of attribute syntax -- i.e. types / terms /
|
wenzelm@16234
|
1077 |
facts mentioned as arguments are always those of the locale definition
|
wenzelm@16234
|
1078 |
context, independently of the context of later invocations. Moreover,
|
wenzelm@16234
|
1079 |
locale operations (renaming and type / term instantiation) are applied
|
wenzelm@16234
|
1080 |
to attribute arguments as expected.
|
wenzelm@16234
|
1081 |
|
wenzelm@16234
|
1082 |
INCOMPATIBILITY of the ML interface: always pass Attrib.src instead of
|
wenzelm@16234
|
1083 |
actual attributes; rare situations may require Attrib.attribute to
|
wenzelm@16234
|
1084 |
embed those attributes into Attrib.src that lack concrete syntax.
|
wenzelm@16234
|
1085 |
Attribute implementations need to cooperate properly with the static
|
wenzelm@16234
|
1086 |
binding mechanism. Basic parsers Args.XXX_typ/term/prop and
|
wenzelm@16234
|
1087 |
Attrib.XXX_thm etc. already do the right thing without further
|
wenzelm@16234
|
1088 |
intervention. Only unusual applications -- such as "where" or "of"
|
wenzelm@16234
|
1089 |
(cf. src/Pure/Isar/attrib.ML), which process arguments depending both
|
wenzelm@16234
|
1090 |
on the context and the facts involved -- may have to assign parsed
|
wenzelm@16234
|
1091 |
values to argument tokens explicitly.
|
wenzelm@16234
|
1092 |
|
wenzelm@16234
|
1093 |
* Changed parameter management in theorem generation for long goal
|
wenzelm@16234
|
1094 |
statements with 'includes'. INCOMPATIBILITY: produces a different
|
wenzelm@16234
|
1095 |
theorem statement in rare situations.
|
wenzelm@16234
|
1096 |
|
ballarin@17228
|
1097 |
* Locale inspection command 'print_locale' omits notes elements. Use
|
ballarin@17228
|
1098 |
'print_locale!' to have them included in the output.
|
ballarin@17228
|
1099 |
|
wenzelm@16234
|
1100 |
|
wenzelm@16234
|
1101 |
*** Provers ***
|
wenzelm@16234
|
1102 |
|
wenzelm@16234
|
1103 |
* Provers/hypsubst.ML: improved version of the subst method, for
|
wenzelm@16234
|
1104 |
single-step rewriting: it now works in bound variable contexts. New is
|
wenzelm@16234
|
1105 |
'subst (asm)', for rewriting an assumption. INCOMPATIBILITY: may
|
wenzelm@16234
|
1106 |
rewrite a different subterm than the original subst method, which is
|
wenzelm@16234
|
1107 |
still available as 'simplesubst'.
|
wenzelm@16234
|
1108 |
|
wenzelm@16234
|
1109 |
* Provers/quasi.ML: new transitivity reasoners for transitivity only
|
wenzelm@16234
|
1110 |
and quasi orders.
|
wenzelm@16234
|
1111 |
|
wenzelm@16234
|
1112 |
* Provers/trancl.ML: new transitivity reasoner for transitive and
|
wenzelm@16234
|
1113 |
reflexive-transitive closure of relations.
|
wenzelm@16234
|
1114 |
|
wenzelm@16234
|
1115 |
* Provers/blast.ML: new reference depth_limit to make blast's depth
|
wenzelm@16234
|
1116 |
limit (previously hard-coded with a value of 20) user-definable.
|
wenzelm@16234
|
1117 |
|
wenzelm@16234
|
1118 |
* Provers/simplifier.ML has been moved to Pure, where Simplifier.setup
|
wenzelm@16234
|
1119 |
is peformed already. Object-logics merely need to finish their
|
wenzelm@16234
|
1120 |
initial simpset configuration as before. INCOMPATIBILITY.
|
wenzelm@15703
|
1121 |
|
berghofe@15475
|
1122 |
|
schirmer@14700
|
1123 |
*** HOL ***
|
schirmer@14700
|
1124 |
|
wenzelm@16234
|
1125 |
* Symbolic syntax of Hilbert Choice Operator is now as follows:
|
wenzelm@14878
|
1126 |
|
wenzelm@14878
|
1127 |
syntax (epsilon)
|
wenzelm@14878
|
1128 |
"_Eps" :: "[pttrn, bool] => 'a" ("(3\<some>_./ _)" [0, 10] 10)
|
wenzelm@14878
|
1129 |
|
wenzelm@16234
|
1130 |
The symbol \<some> is displayed as the alternative epsilon of LaTeX
|
wenzelm@16234
|
1131 |
and x-symbol; use option '-m epsilon' to get it actually printed.
|
wenzelm@16234
|
1132 |
Moreover, the mathematically important symbolic identifier \<epsilon>
|
wenzelm@16234
|
1133 |
becomes available as variable, constant etc. INCOMPATIBILITY,
|
wenzelm@16234
|
1134 |
|
wenzelm@16234
|
1135 |
* "x > y" abbreviates "y < x" and "x >= y" abbreviates "y <= x".
|
wenzelm@16234
|
1136 |
Similarly for all quantifiers: "ALL x > y" etc. The x-symbol for >=
|
wenzelm@17371
|
1137 |
is \<ge>. New transitivity rules have been added to HOL/Orderings.thy to
|
avigad@17016
|
1138 |
support corresponding Isar calculations.
|
wenzelm@16234
|
1139 |
|
wenzelm@16234
|
1140 |
* "{x:A. P}" abbreviates "{x. x:A & P}", and similarly for "\<in>"
|
wenzelm@16234
|
1141 |
instead of ":".
|
wenzelm@16234
|
1142 |
|
wenzelm@16234
|
1143 |
* theory SetInterval: changed the syntax for open intervals:
|
wenzelm@16234
|
1144 |
|
wenzelm@16234
|
1145 |
Old New
|
wenzelm@16234
|
1146 |
{..n(} {..<n}
|
wenzelm@16234
|
1147 |
{)n..} {n<..}
|
wenzelm@16234
|
1148 |
{m..n(} {m..<n}
|
wenzelm@16234
|
1149 |
{)m..n} {m<..n}
|
wenzelm@16234
|
1150 |
{)m..n(} {m<..<n}
|
wenzelm@16234
|
1151 |
|
wenzelm@16234
|
1152 |
The old syntax is still supported but will disappear in the next
|
wenzelm@16234
|
1153 |
release. For conversion use the following Emacs search and replace
|
wenzelm@16234
|
1154 |
patterns (these are not perfect but work quite well):
|
nipkow@15046
|
1155 |
|
nipkow@15046
|
1156 |
{)\([^\.]*\)\.\. -> {\1<\.\.}
|
nipkow@15046
|
1157 |
\.\.\([^(}]*\)(} -> \.\.<\1}
|
nipkow@15046
|
1158 |
|
wenzelm@17533
|
1159 |
* Theory Commutative_Ring (in Library): method comm_ring for proving
|
wenzelm@17533
|
1160 |
equalities in commutative rings; method 'algebra' provides a generic
|
wenzelm@17533
|
1161 |
interface.
|
wenzelm@17389
|
1162 |
|
wenzelm@17389
|
1163 |
* Theory Finite_Set: changed the syntax for 'setsum', summation over
|
wenzelm@16234
|
1164 |
finite sets: "setsum (%x. e) A", which used to be "\<Sum>x:A. e", is
|
wenzelm@17371
|
1165 |
now either "SUM x:A. e" or "\<Sum>x \<in> A. e". The bound variable can
|
paulson@17189
|
1166 |
be a tuple pattern.
|
wenzelm@16234
|
1167 |
|
wenzelm@16234
|
1168 |
Some new syntax forms are available:
|
wenzelm@16234
|
1169 |
|
wenzelm@16234
|
1170 |
"\<Sum>x | P. e" for "setsum (%x. e) {x. P}"
|
wenzelm@16234
|
1171 |
"\<Sum>x = a..b. e" for "setsum (%x. e) {a..b}"
|
wenzelm@16234
|
1172 |
"\<Sum>x = a..<b. e" for "setsum (%x. e) {a..<b}"
|
wenzelm@16234
|
1173 |
"\<Sum>x < k. e" for "setsum (%x. e) {..<k}"
|
wenzelm@16234
|
1174 |
|
wenzelm@16234
|
1175 |
The latter form "\<Sum>x < k. e" used to be based on a separate
|
wenzelm@16234
|
1176 |
function "Summation", which has been discontinued.
|
wenzelm@16234
|
1177 |
|
wenzelm@16234
|
1178 |
* theory Finite_Set: in structured induction proofs, the insert case
|
wenzelm@16234
|
1179 |
is now 'case (insert x F)' instead of the old counterintuitive 'case
|
wenzelm@16234
|
1180 |
(insert F x)'.
|
wenzelm@16234
|
1181 |
|
wenzelm@16234
|
1182 |
* The 'refute' command has been extended to support a much larger
|
wenzelm@16234
|
1183 |
fragment of HOL, including axiomatic type classes, constdefs and
|
wenzelm@16234
|
1184 |
typedefs, inductive datatypes and recursion.
|
wenzelm@16234
|
1185 |
|
webertj@17700
|
1186 |
* New tactics 'sat' and 'satx' to prove propositional tautologies.
|
webertj@17700
|
1187 |
Requires zChaff with proof generation to be installed. See
|
webertj@17700
|
1188 |
HOL/ex/SAT_Examples.thy for examples.
|
webertj@17619
|
1189 |
|
wenzelm@16234
|
1190 |
* Datatype induction via method 'induct' now preserves the name of the
|
wenzelm@16234
|
1191 |
induction variable. For example, when proving P(xs::'a list) by
|
wenzelm@16234
|
1192 |
induction on xs, the induction step is now P(xs) ==> P(a#xs) rather
|
wenzelm@16234
|
1193 |
than P(list) ==> P(a#list) as previously. Potential INCOMPATIBILITY
|
wenzelm@16234
|
1194 |
in unstructured proof scripts.
|
wenzelm@16234
|
1195 |
|
wenzelm@16234
|
1196 |
* Reworked implementation of records. Improved scalability for
|
wenzelm@16234
|
1197 |
records with many fields, avoiding performance problems for type
|
wenzelm@16234
|
1198 |
inference. Records are no longer composed of nested field types, but
|
wenzelm@16234
|
1199 |
of nested extension types. Therefore the record type only grows linear
|
wenzelm@16234
|
1200 |
in the number of extensions and not in the number of fields. The
|
wenzelm@16234
|
1201 |
top-level (users) view on records is preserved. Potential
|
wenzelm@16234
|
1202 |
INCOMPATIBILITY only in strange cases, where the theory depends on the
|
wenzelm@16234
|
1203 |
old record representation. The type generated for a record is called
|
wenzelm@16234
|
1204 |
<record_name>_ext_type.
|
wenzelm@16234
|
1205 |
|
wenzelm@16234
|
1206 |
Flag record_quick_and_dirty_sensitive can be enabled to skip the
|
wenzelm@16234
|
1207 |
proofs triggered by a record definition or a simproc (if
|
wenzelm@16234
|
1208 |
quick_and_dirty is enabled). Definitions of large records can take
|
wenzelm@16234
|
1209 |
quite long.
|
wenzelm@16234
|
1210 |
|
wenzelm@16234
|
1211 |
New simproc record_upd_simproc for simplification of multiple record
|
wenzelm@16234
|
1212 |
updates enabled by default. Moreover, trivial updates are also
|
wenzelm@16234
|
1213 |
removed: r(|x := x r|) = r. INCOMPATIBILITY: old proofs break
|
wenzelm@16234
|
1214 |
occasionally, since simplification is more powerful by default.
|
wenzelm@16234
|
1215 |
|
wenzelm@17275
|
1216 |
* typedef: proper support for polymorphic sets, which contain extra
|
wenzelm@17275
|
1217 |
type-variables in the term.
|
wenzelm@17275
|
1218 |
|
wenzelm@16234
|
1219 |
* Simplifier: automatically reasons about transitivity chains
|
wenzelm@16234
|
1220 |
involving "trancl" (r^+) and "rtrancl" (r^*) by setting up tactics
|
wenzelm@16234
|
1221 |
provided by Provers/trancl.ML as additional solvers. INCOMPATIBILITY:
|
wenzelm@16234
|
1222 |
old proofs break occasionally as simplification may now solve more
|
wenzelm@16234
|
1223 |
goals than previously.
|
wenzelm@16234
|
1224 |
|
wenzelm@16234
|
1225 |
* Simplifier: converts x <= y into x = y if assumption y <= x is
|
wenzelm@16234
|
1226 |
present. Works for all partial orders (class "order"), in particular
|
wenzelm@16234
|
1227 |
numbers and sets. For linear orders (e.g. numbers) it treats ~ x < y
|
wenzelm@16234
|
1228 |
just like y <= x.
|
wenzelm@16234
|
1229 |
|
wenzelm@16234
|
1230 |
* Simplifier: new simproc for "let x = a in f x". If a is a free or
|
wenzelm@16234
|
1231 |
bound variable or a constant then the let is unfolded. Otherwise
|
wenzelm@16234
|
1232 |
first a is simplified to b, and then f b is simplified to g. If
|
wenzelm@16234
|
1233 |
possible we abstract b from g arriving at "let x = b in h x",
|
wenzelm@16234
|
1234 |
otherwise we unfold the let and arrive at g. The simproc can be
|
wenzelm@16234
|
1235 |
enabled/disabled by the reference use_let_simproc. Potential
|
wenzelm@16234
|
1236 |
INCOMPATIBILITY since simplification is more powerful by default.
|
webertj@15776
|
1237 |
|
paulson@16563
|
1238 |
* Classical reasoning: the meson method now accepts theorems as arguments.
|
paulson@16563
|
1239 |
|
paulson@17595
|
1240 |
* Prover support: pre-release of the Isabelle-ATP linkup, which runs background
|
paulson@17595
|
1241 |
jobs to provide advice on the provability of subgoals.
|
paulson@17595
|
1242 |
|
wenzelm@16891
|
1243 |
* Theory OrderedGroup and Ring_and_Field: various additions and
|
wenzelm@16891
|
1244 |
improvements to faciliate calculations involving equalities and
|
wenzelm@16891
|
1245 |
inequalities.
|
wenzelm@16891
|
1246 |
|
wenzelm@16891
|
1247 |
The following theorems have been eliminated or modified
|
wenzelm@16891
|
1248 |
(INCOMPATIBILITY):
|
avigad@16888
|
1249 |
|
avigad@16888
|
1250 |
abs_eq now named abs_of_nonneg
|
wenzelm@17371
|
1251 |
abs_of_ge_0 now named abs_of_nonneg
|
wenzelm@17371
|
1252 |
abs_minus_eq now named abs_of_nonpos
|
avigad@16888
|
1253 |
imp_abs_id now named abs_of_nonneg
|
avigad@16888
|
1254 |
imp_abs_neg_id now named abs_of_nonpos
|
avigad@16888
|
1255 |
mult_pos now named mult_pos_pos
|
avigad@16888
|
1256 |
mult_pos_le now named mult_nonneg_nonneg
|
avigad@16888
|
1257 |
mult_pos_neg_le now named mult_nonneg_nonpos
|
avigad@16888
|
1258 |
mult_pos_neg2_le now named mult_nonneg_nonpos2
|
avigad@16888
|
1259 |
mult_neg now named mult_neg_neg
|
avigad@16888
|
1260 |
mult_neg_le now named mult_nonpos_nonpos
|
avigad@16888
|
1261 |
|
wenzelm@16891
|
1262 |
* Theory Parity: added rules for simplifying exponents.
|
wenzelm@16891
|
1263 |
|
nipkow@17092
|
1264 |
* Theory List:
|
nipkow@17092
|
1265 |
|
nipkow@17092
|
1266 |
The following theorems have been eliminated or modified
|
nipkow@17092
|
1267 |
(INCOMPATIBILITY):
|
nipkow@17092
|
1268 |
|
nipkow@17092
|
1269 |
list_all_Nil now named list_all.simps(1)
|
nipkow@17092
|
1270 |
list_all_Cons now named list_all.simps(2)
|
nipkow@17092
|
1271 |
list_all_conv now named list_all_iff
|
nipkow@17092
|
1272 |
set_mem_eq now named mem_iff
|
nipkow@17092
|
1273 |
|
wenzelm@16929
|
1274 |
* Theories SetsAndFunctions and BigO (see HOL/Library) support
|
wenzelm@16929
|
1275 |
asymptotic "big O" calculations. See the notes in BigO.thy.
|
wenzelm@16929
|
1276 |
|
avigad@16888
|
1277 |
|
avigad@16888
|
1278 |
*** HOL-Complex ***
|
avigad@16888
|
1279 |
|
wenzelm@16891
|
1280 |
* Theory RealDef: better support for embedding natural numbers and
|
wenzelm@16891
|
1281 |
integers in the reals.
|
wenzelm@16891
|
1282 |
|
wenzelm@16891
|
1283 |
The following theorems have been eliminated or modified
|
wenzelm@16891
|
1284 |
(INCOMPATIBILITY):
|
wenzelm@16891
|
1285 |
|
avigad@17016
|
1286 |
exp_ge_add_one_self now requires no hypotheses
|
avigad@17016
|
1287 |
real_of_int_add reversed direction of equality (use [symmetric])
|
avigad@17016
|
1288 |
real_of_int_minus reversed direction of equality (use [symmetric])
|
avigad@17016
|
1289 |
real_of_int_diff reversed direction of equality (use [symmetric])
|
avigad@17016
|
1290 |
real_of_int_mult reversed direction of equality (use [symmetric])
|
wenzelm@16891
|
1291 |
|
wenzelm@16891
|
1292 |
* Theory RComplete: expanded support for floor and ceiling functions.
|
avigad@16888
|
1293 |
|
avigad@16962
|
1294 |
* Theory Ln is new, with properties of the natural logarithm
|
avigad@16962
|
1295 |
|
wenzelm@17423
|
1296 |
* Hyperreal: There is a new type constructor "star" for making
|
wenzelm@17423
|
1297 |
nonstandard types. The old type names are now type synonyms:
|
wenzelm@17423
|
1298 |
|
wenzelm@17423
|
1299 |
hypreal = real star
|
wenzelm@17423
|
1300 |
hypnat = nat star
|
wenzelm@17423
|
1301 |
hcomplex = complex star
|
wenzelm@17423
|
1302 |
|
wenzelm@17423
|
1303 |
* Hyperreal: Many groups of similarly-defined constants have been
|
huffman@17442
|
1304 |
replaced by polymorphic versions (INCOMPATIBILITY):
|
wenzelm@17423
|
1305 |
|
wenzelm@17423
|
1306 |
star_of <-- hypreal_of_real, hypnat_of_nat, hcomplex_of_complex
|
wenzelm@17423
|
1307 |
|
wenzelm@17423
|
1308 |
starset <-- starsetNat, starsetC
|
wenzelm@17423
|
1309 |
*s* <-- *sNat*, *sc*
|
wenzelm@17423
|
1310 |
starset_n <-- starsetNat_n, starsetC_n
|
wenzelm@17423
|
1311 |
*sn* <-- *sNatn*, *scn*
|
wenzelm@17423
|
1312 |
InternalSets <-- InternalNatSets, InternalCSets
|
wenzelm@17423
|
1313 |
|
huffman@17442
|
1314 |
starfun <-- starfun{Nat,Nat2,C,RC,CR}
|
wenzelm@17423
|
1315 |
*f* <-- *fNat*, *fNat2*, *fc*, *fRc*, *fcR*
|
huffman@17442
|
1316 |
starfun_n <-- starfun{Nat,Nat2,C,RC,CR}_n
|
wenzelm@17423
|
1317 |
*fn* <-- *fNatn*, *fNat2n*, *fcn*, *fRcn*, *fcRn*
|
huffman@17442
|
1318 |
InternalFuns <-- InternalNatFuns, InternalNatFuns2, Internal{C,RC,CR}Funs
|
wenzelm@17423
|
1319 |
|
wenzelm@17423
|
1320 |
* Hyperreal: Many type-specific theorems have been removed in favor of
|
huffman@17442
|
1321 |
theorems specific to various axiomatic type classes (INCOMPATIBILITY):
|
huffman@17442
|
1322 |
|
huffman@17442
|
1323 |
add_commute <-- {hypreal,hypnat,hcomplex}_add_commute
|
huffman@17442
|
1324 |
add_assoc <-- {hypreal,hypnat,hcomplex}_add_assocs
|
huffman@17442
|
1325 |
OrderedGroup.add_0 <-- {hypreal,hypnat,hcomplex}_add_zero_left
|
huffman@17442
|
1326 |
OrderedGroup.add_0_right <-- {hypreal,hcomplex}_add_zero_right
|
wenzelm@17423
|
1327 |
right_minus <-- hypreal_add_minus
|
huffman@17442
|
1328 |
left_minus <-- {hypreal,hcomplex}_add_minus_left
|
huffman@17442
|
1329 |
mult_commute <-- {hypreal,hypnat,hcomplex}_mult_commute
|
huffman@17442
|
1330 |
mult_assoc <-- {hypreal,hypnat,hcomplex}_mult_assoc
|
huffman@17442
|
1331 |
mult_1_left <-- {hypreal,hypnat}_mult_1, hcomplex_mult_one_left
|
wenzelm@17423
|
1332 |
mult_1_right <-- hcomplex_mult_one_right
|
wenzelm@17423
|
1333 |
mult_zero_left <-- hcomplex_mult_zero_left
|
huffman@17442
|
1334 |
left_distrib <-- {hypreal,hypnat,hcomplex}_add_mult_distrib
|
wenzelm@17423
|
1335 |
right_distrib <-- hypnat_add_mult_distrib2
|
huffman@17442
|
1336 |
zero_neq_one <-- {hypreal,hypnat,hcomplex}_zero_not_eq_one
|
wenzelm@17423
|
1337 |
right_inverse <-- hypreal_mult_inverse
|
wenzelm@17423
|
1338 |
left_inverse <-- hypreal_mult_inverse_left, hcomplex_mult_inv_left
|
huffman@17442
|
1339 |
order_refl <-- {hypreal,hypnat}_le_refl
|
huffman@17442
|
1340 |
order_trans <-- {hypreal,hypnat}_le_trans
|
huffman@17442
|
1341 |
order_antisym <-- {hypreal,hypnat}_le_anti_sym
|
huffman@17442
|
1342 |
order_less_le <-- {hypreal,hypnat}_less_le
|
huffman@17442
|
1343 |
linorder_linear <-- {hypreal,hypnat}_le_linear
|
huffman@17442
|
1344 |
add_left_mono <-- {hypreal,hypnat}_add_left_mono
|
huffman@17442
|
1345 |
mult_strict_left_mono <-- {hypreal,hypnat}_mult_less_mono2
|
wenzelm@17423
|
1346 |
add_nonneg_nonneg <-- hypreal_le_add_order
|
wenzelm@17423
|
1347 |
|
wenzelm@17423
|
1348 |
* Hyperreal: Separate theorems having to do with type-specific
|
wenzelm@17423
|
1349 |
versions of constants have been merged into theorems that apply to the
|
huffman@17442
|
1350 |
new polymorphic constants (INCOMPATIBILITY):
|
huffman@17442
|
1351 |
|
huffman@17442
|
1352 |
STAR_UNIV_set <-- {STAR_real,NatStar_real,STARC_complex}_set
|
huffman@17442
|
1353 |
STAR_empty_set <-- {STAR,NatStar,STARC}_empty_set
|
huffman@17442
|
1354 |
STAR_Un <-- {STAR,NatStar,STARC}_Un
|
huffman@17442
|
1355 |
STAR_Int <-- {STAR,NatStar,STARC}_Int
|
huffman@17442
|
1356 |
STAR_Compl <-- {STAR,NatStar,STARC}_Compl
|
huffman@17442
|
1357 |
STAR_subset <-- {STAR,NatStar,STARC}_subset
|
huffman@17442
|
1358 |
STAR_mem <-- {STAR,NatStar,STARC}_mem
|
huffman@17442
|
1359 |
STAR_mem_Compl <-- {STAR,STARC}_mem_Compl
|
huffman@17442
|
1360 |
STAR_diff <-- {STAR,STARC}_diff
|
huffman@17442
|
1361 |
STAR_star_of_image_subset <-- {STAR_hypreal_of_real, NatStar_hypreal_of_real,
|
huffman@17442
|
1362 |
STARC_hcomplex_of_complex}_image_subset
|
huffman@17442
|
1363 |
starset_n_Un <-- starset{Nat,C}_n_Un
|
huffman@17442
|
1364 |
starset_n_Int <-- starset{Nat,C}_n_Int
|
huffman@17442
|
1365 |
starset_n_Compl <-- starset{Nat,C}_n_Compl
|
huffman@17442
|
1366 |
starset_n_diff <-- starset{Nat,C}_n_diff
|
huffman@17442
|
1367 |
InternalSets_Un <-- Internal{Nat,C}Sets_Un
|
huffman@17442
|
1368 |
InternalSets_Int <-- Internal{Nat,C}Sets_Int
|
huffman@17442
|
1369 |
InternalSets_Compl <-- Internal{Nat,C}Sets_Compl
|
huffman@17442
|
1370 |
InternalSets_diff <-- Internal{Nat,C}Sets_diff
|
huffman@17442
|
1371 |
InternalSets_UNIV_diff <-- Internal{Nat,C}Sets_UNIV_diff
|
huffman@17442
|
1372 |
InternalSets_starset_n <-- Internal{Nat,C}Sets_starset{Nat,C}_n
|
huffman@17442
|
1373 |
starset_starset_n_eq <-- starset{Nat,C}_starset{Nat,C}_n_eq
|
huffman@17442
|
1374 |
starset_n_starset <-- starset{Nat,C}_n_starset{Nat,C}
|
huffman@17442
|
1375 |
starfun_n_starfun <-- starfun{Nat,Nat2,C,RC,CR}_n_starfun{Nat,Nat2,C,RC,CR}
|
huffman@17442
|
1376 |
starfun <-- starfun{Nat,Nat2,C,RC,CR}
|
huffman@17442
|
1377 |
starfun_mult <-- starfun{Nat,Nat2,C,RC,CR}_mult
|
huffman@17442
|
1378 |
starfun_add <-- starfun{Nat,Nat2,C,RC,CR}_add
|
huffman@17442
|
1379 |
starfun_minus <-- starfun{Nat,Nat2,C,RC,CR}_minus
|
huffman@17442
|
1380 |
starfun_diff <-- starfun{C,RC,CR}_diff
|
huffman@17442
|
1381 |
starfun_o <-- starfun{NatNat2,Nat2,_stafunNat,C,C_starfunRC,_starfunCR}_o
|
huffman@17442
|
1382 |
starfun_o2 <-- starfun{NatNat2,_stafunNat,C,C_starfunRC,_starfunCR}_o2
|
huffman@17442
|
1383 |
starfun_const_fun <-- starfun{Nat,Nat2,C,RC,CR}_const_fun
|
huffman@17442
|
1384 |
starfun_inverse <-- starfun{Nat,C,RC,CR}_inverse
|
huffman@17442
|
1385 |
starfun_eq <-- starfun{Nat,Nat2,C,RC,CR}_eq
|
huffman@17442
|
1386 |
starfun_eq_iff <-- starfun{C,RC,CR}_eq_iff
|
wenzelm@17423
|
1387 |
starfun_Id <-- starfunC_Id
|
huffman@17442
|
1388 |
starfun_approx <-- starfun{Nat,CR}_approx
|
huffman@17442
|
1389 |
starfun_capprox <-- starfun{C,RC}_capprox
|
wenzelm@17423
|
1390 |
starfun_abs <-- starfunNat_rabs
|
huffman@17442
|
1391 |
starfun_lambda_cancel <-- starfun{C,CR,RC}_lambda_cancel
|
huffman@17442
|
1392 |
starfun_lambda_cancel2 <-- starfun{C,CR,RC}_lambda_cancel2
|
wenzelm@17423
|
1393 |
starfun_mult_HFinite_approx <-- starfunCR_mult_HFinite_capprox
|
huffman@17442
|
1394 |
starfun_mult_CFinite_capprox <-- starfun{C,RC}_mult_CFinite_capprox
|
huffman@17442
|
1395 |
starfun_add_capprox <-- starfun{C,RC}_add_capprox
|
wenzelm@17423
|
1396 |
starfun_add_approx <-- starfunCR_add_approx
|
wenzelm@17423
|
1397 |
starfun_inverse_inverse <-- starfunC_inverse_inverse
|
huffman@17442
|
1398 |
starfun_divide <-- starfun{C,CR,RC}_divide
|
huffman@17442
|
1399 |
starfun_n <-- starfun{Nat,C}_n
|
huffman@17442
|
1400 |
starfun_n_mult <-- starfun{Nat,C}_n_mult
|
huffman@17442
|
1401 |
starfun_n_add <-- starfun{Nat,C}_n_add
|
wenzelm@17423
|
1402 |
starfun_n_add_minus <-- starfunNat_n_add_minus
|
huffman@17442
|
1403 |
starfun_n_const_fun <-- starfun{Nat,C}_n_const_fun
|
huffman@17442
|
1404 |
starfun_n_minus <-- starfun{Nat,C}_n_minus
|
huffman@17442
|
1405 |
starfun_n_eq <-- starfun{Nat,C}_n_eq
|
huffman@17442
|
1406 |
|
huffman@17442
|
1407 |
star_n_add <-- {hypreal,hypnat,hcomplex}_add
|
huffman@17442
|
1408 |
star_n_minus <-- {hypreal,hcomplex}_minus
|
huffman@17442
|
1409 |
star_n_diff <-- {hypreal,hcomplex}_diff
|
huffman@17442
|
1410 |
star_n_mult <-- {hypreal,hcomplex}_mult
|
huffman@17442
|
1411 |
star_n_inverse <-- {hypreal,hcomplex}_inverse
|
huffman@17442
|
1412 |
star_n_le <-- {hypreal,hypnat}_le
|
huffman@17442
|
1413 |
star_n_less <-- {hypreal,hypnat}_less
|
huffman@17442
|
1414 |
star_n_zero_num <-- {hypreal,hypnat,hcomplex}_zero_num
|
huffman@17442
|
1415 |
star_n_one_num <-- {hypreal,hypnat,hcomplex}_one_num
|
wenzelm@17423
|
1416 |
star_n_abs <-- hypreal_hrabs
|
wenzelm@17423
|
1417 |
star_n_divide <-- hcomplex_divide
|
wenzelm@17423
|
1418 |
|
huffman@17442
|
1419 |
star_of_add <-- {hypreal_of_real,hypnat_of_nat,hcomplex_of_complex}_add
|
huffman@17442
|
1420 |
star_of_minus <-- {hypreal_of_real,hcomplex_of_complex}_minus
|
wenzelm@17423
|
1421 |
star_of_diff <-- hypreal_of_real_diff
|
huffman@17442
|
1422 |
star_of_mult <-- {hypreal_of_real,hypnat_of_nat,hcomplex_of_complex}_mult
|
huffman@17442
|
1423 |
star_of_one <-- {hypreal_of_real,hcomplex_of_complex}_one
|
huffman@17442
|
1424 |
star_of_zero <-- {hypreal_of_real,hypnat_of_nat,hcomplex_of_complex}_zero
|
huffman@17442
|
1425 |
star_of_le <-- {hypreal_of_real,hypnat_of_nat}_le_iff
|
huffman@17442
|
1426 |
star_of_less <-- {hypreal_of_real,hypnat_of_nat}_less_iff
|
huffman@17442
|
1427 |
star_of_eq <-- {hypreal_of_real,hypnat_of_nat,hcomplex_of_complex}_eq_iff
|
huffman@17442
|
1428 |
star_of_inverse <-- {hypreal_of_real,hcomplex_of_complex}_inverse
|
huffman@17442
|
1429 |
star_of_divide <-- {hypreal_of_real,hcomplex_of_complex}_divide
|
huffman@17442
|
1430 |
star_of_of_nat <-- {hypreal_of_real,hcomplex_of_complex}_of_nat
|
huffman@17442
|
1431 |
star_of_of_int <-- {hypreal_of_real,hcomplex_of_complex}_of_int
|
huffman@17442
|
1432 |
star_of_number_of <-- {hypreal,hcomplex}_number_of
|
wenzelm@17423
|
1433 |
star_of_number_less <-- number_of_less_hypreal_of_real_iff
|
wenzelm@17423
|
1434 |
star_of_number_le <-- number_of_le_hypreal_of_real_iff
|
wenzelm@17423
|
1435 |
star_of_eq_number <-- hypreal_of_real_eq_number_of_iff
|
wenzelm@17423
|
1436 |
star_of_less_number <-- hypreal_of_real_less_number_of_iff
|
wenzelm@17423
|
1437 |
star_of_le_number <-- hypreal_of_real_le_number_of_iff
|
wenzelm@17423
|
1438 |
star_of_power <-- hypreal_of_real_power
|
wenzelm@17423
|
1439 |
star_of_eq_0 <-- hcomplex_of_complex_zero_iff
|
wenzelm@17423
|
1440 |
|
huffman@17442
|
1441 |
* Hyperreal: new method "transfer" that implements the transfer
|
huffman@17442
|
1442 |
principle of nonstandard analysis. With a subgoal that mentions
|
huffman@17442
|
1443 |
nonstandard types like "'a star", the command "apply transfer"
|
huffman@17442
|
1444 |
replaces it with an equivalent one that mentions only standard types.
|
huffman@17442
|
1445 |
To be successful, all free variables must have standard types; non-
|
huffman@17442
|
1446 |
standard variables must have explicit universal quantifiers.
|
huffman@17442
|
1447 |
|
wenzelm@17641
|
1448 |
* Hyperreal: A theory of Taylor series.
|
wenzelm@17641
|
1449 |
|
wenzelm@14655
|
1450 |
|
wenzelm@14682
|
1451 |
*** HOLCF ***
|
wenzelm@14682
|
1452 |
|
wenzelm@17533
|
1453 |
* Discontinued special version of 'constdefs' (which used to support
|
wenzelm@17533
|
1454 |
continuous functions) in favor of the general Pure one with full
|
wenzelm@17533
|
1455 |
type-inference.
|
wenzelm@17533
|
1456 |
|
wenzelm@17533
|
1457 |
* New simplification procedure for solving continuity conditions; it
|
wenzelm@17533
|
1458 |
is much faster on terms with many nested lambda abstractions (cubic
|
huffman@17442
|
1459 |
instead of exponential time).
|
huffman@17442
|
1460 |
|
wenzelm@17533
|
1461 |
* New syntax for domain package: selector names are now optional.
|
huffman@17442
|
1462 |
Parentheses should be omitted unless argument is lazy, for example:
|
huffman@17442
|
1463 |
|
huffman@17442
|
1464 |
domain 'a stream = cons "'a" (lazy "'a stream")
|
huffman@17442
|
1465 |
|
wenzelm@17533
|
1466 |
* New command 'fixrec' for defining recursive functions with pattern
|
wenzelm@17533
|
1467 |
matching; defining multiple functions with mutual recursion is also
|
wenzelm@17533
|
1468 |
supported. Patterns may include the constants cpair, spair, up, sinl,
|
wenzelm@17533
|
1469 |
sinr, or any data constructor defined by the domain package. The given
|
wenzelm@17533
|
1470 |
equations are proven as rewrite rules. See HOLCF/ex/Fixrec_ex.thy for
|
wenzelm@17533
|
1471 |
syntax and examples.
|
wenzelm@17533
|
1472 |
|
wenzelm@17533
|
1473 |
* New commands 'cpodef' and 'pcpodef' for defining predicate subtypes
|
wenzelm@17533
|
1474 |
of cpo and pcpo types. Syntax is exactly like the 'typedef' command,
|
wenzelm@17533
|
1475 |
but the proof obligation additionally includes an admissibility
|
wenzelm@17533
|
1476 |
requirement. The packages generate instances of class cpo or pcpo,
|
wenzelm@17533
|
1477 |
with continuity and strictness theorems for Rep and Abs.
|
huffman@17442
|
1478 |
|
huffman@17584
|
1479 |
* HOLCF: Many theorems have been renamed according to a more standard naming
|
huffman@17584
|
1480 |
scheme (INCOMPATIBILITY):
|
huffman@17584
|
1481 |
|
huffman@17584
|
1482 |
foo_inject: "foo$x = foo$y ==> x = y"
|
huffman@17584
|
1483 |
foo_eq: "(foo$x = foo$y) = (x = y)"
|
huffman@17584
|
1484 |
foo_less: "(foo$x << foo$y) = (x << y)"
|
huffman@17584
|
1485 |
foo_strict: "foo$UU = UU"
|
huffman@17584
|
1486 |
foo_defined: "... ==> foo$x ~= UU"
|
huffman@17584
|
1487 |
foo_defined_iff: "(foo$x = UU) = (x = UU)"
|
huffman@17584
|
1488 |
|
wenzelm@14682
|
1489 |
|
paulson@14885
|
1490 |
*** ZF ***
|
paulson@14885
|
1491 |
|
wenzelm@16234
|
1492 |
* ZF/ex: theories Group and Ring provide examples in abstract algebra,
|
wenzelm@16234
|
1493 |
including the First Isomorphism Theorem (on quotienting by the kernel
|
wenzelm@16234
|
1494 |
of a homomorphism).
|
wenzelm@15089
|
1495 |
|
wenzelm@15089
|
1496 |
* ZF/Simplifier: install second copy of type solver that actually
|
wenzelm@16234
|
1497 |
makes use of TC rules declared to Isar proof contexts (or locales);
|
wenzelm@16234
|
1498 |
the old version is still required for ML proof scripts.
|
wenzelm@15703
|
1499 |
|
wenzelm@15703
|
1500 |
|
wenzelm@17445
|
1501 |
*** Cube ***
|
wenzelm@17445
|
1502 |
|
wenzelm@17445
|
1503 |
* Converted to Isar theory format; use locales instead of axiomatic
|
wenzelm@17445
|
1504 |
theories.
|
wenzelm@17445
|
1505 |
|
wenzelm@17445
|
1506 |
|
wenzelm@15703
|
1507 |
*** ML ***
|
wenzelm@15703
|
1508 |
|
wenzelm@15973
|
1509 |
* Pure/library.ML no longer defines its own option datatype, but uses
|
wenzelm@16234
|
1510 |
that of the SML basis, which has constructors NONE and SOME instead of
|
wenzelm@16234
|
1511 |
None and Some, as well as exception Option.Option instead of OPTION.
|
wenzelm@16234
|
1512 |
The functions the, if_none, is_some, is_none have been adapted
|
wenzelm@16234
|
1513 |
accordingly, while Option.map replaces apsome.
|
wenzelm@15973
|
1514 |
|
wenzelm@16860
|
1515 |
* Pure/library.ML: the exception LIST has been given up in favour of
|
wenzelm@16860
|
1516 |
the standard exceptions Empty and Subscript, as well as
|
wenzelm@16860
|
1517 |
Library.UnequalLengths. Function like Library.hd and Library.tl are
|
wenzelm@16860
|
1518 |
superceded by the standard hd and tl functions etc.
|
wenzelm@16860
|
1519 |
|
wenzelm@16860
|
1520 |
A number of basic list functions are no longer exported to the ML
|
wenzelm@16860
|
1521 |
toplevel, as they are variants of predefined functions. The following
|
wenzelm@16234
|
1522 |
suggests how one can translate existing code:
|
wenzelm@15973
|
1523 |
|
wenzelm@15973
|
1524 |
rev_append xs ys = List.revAppend (xs, ys)
|
wenzelm@15973
|
1525 |
nth_elem (i, xs) = List.nth (xs, i)
|
wenzelm@15973
|
1526 |
last_elem xs = List.last xs
|
wenzelm@15973
|
1527 |
flat xss = List.concat xss
|
wenzelm@16234
|
1528 |
seq fs = List.app fs
|
wenzelm@15973
|
1529 |
partition P xs = List.partition P xs
|
wenzelm@15973
|
1530 |
mapfilter f xs = List.mapPartial f xs
|
wenzelm@15973
|
1531 |
|
wenzelm@16860
|
1532 |
* Pure/library.ML: several combinators for linear functional
|
wenzelm@16860
|
1533 |
transformations, notably reverse application and composition:
|
wenzelm@16860
|
1534 |
|
wenzelm@17371
|
1535 |
x |> f f #> g
|
wenzelm@17371
|
1536 |
(x, y) |-> f f #-> g
|
wenzelm@16860
|
1537 |
|
haftmann@17495
|
1538 |
* Pure/library.ML: introduced/changed precedence of infix operators:
|
haftmann@17495
|
1539 |
|
haftmann@17495
|
1540 |
infix 1 |> |-> ||> ||>> |>> |>>> #> #->;
|
haftmann@17495
|
1541 |
infix 2 ?;
|
haftmann@17495
|
1542 |
infix 3 o oo ooo oooo;
|
haftmann@17495
|
1543 |
infix 4 ~~ upto downto;
|
haftmann@17495
|
1544 |
|
haftmann@17495
|
1545 |
Maybe INCOMPATIBILITY when any of those is used in conjunction with other
|
haftmann@17495
|
1546 |
infix operators.
|
haftmann@17495
|
1547 |
|
wenzelm@17408
|
1548 |
* Pure/library.ML: natural list combinators fold, fold_rev, and
|
haftmann@16869
|
1549 |
fold_map support linear functional transformations and nesting. For
|
wenzelm@16860
|
1550 |
example:
|
wenzelm@16860
|
1551 |
|
wenzelm@16860
|
1552 |
fold f [x1, ..., xN] y =
|
wenzelm@16860
|
1553 |
y |> f x1 |> ... |> f xN
|
wenzelm@16860
|
1554 |
|
wenzelm@16860
|
1555 |
(fold o fold) f [xs1, ..., xsN] y =
|
wenzelm@16860
|
1556 |
y |> fold f xs1 |> ... |> fold f xsN
|
wenzelm@16860
|
1557 |
|
wenzelm@16860
|
1558 |
fold f [x1, ..., xN] =
|
wenzelm@16860
|
1559 |
f x1 #> ... #> f xN
|
wenzelm@16860
|
1560 |
|
wenzelm@16860
|
1561 |
(fold o fold) f [xs1, ..., xsN] =
|
wenzelm@16860
|
1562 |
fold f xs1 #> ... #> fold f xsN
|
wenzelm@16860
|
1563 |
|
wenzelm@17408
|
1564 |
* Pure/library.ML: the following selectors on type 'a option are
|
wenzelm@17408
|
1565 |
available:
|
wenzelm@17408
|
1566 |
|
wenzelm@17408
|
1567 |
the: 'a option -> 'a (*partial*)
|
wenzelm@17408
|
1568 |
these: 'a option -> 'a where 'a = 'b list
|
haftmann@17402
|
1569 |
the_default: 'a -> 'a option -> 'a
|
haftmann@17402
|
1570 |
the_list: 'a option -> 'a list
|
haftmann@17402
|
1571 |
|
wenzelm@17408
|
1572 |
* Pure/General: structure AList (cf. Pure/General/alist.ML) provides
|
wenzelm@17408
|
1573 |
basic operations for association lists, following natural argument
|
haftmann@17564
|
1574 |
order; moreover the explicit equality predicate passed here avoids
|
haftmann@17495
|
1575 |
potentially expensive polymorphic runtime equality checks.
|
haftmann@17495
|
1576 |
The old functions may be expressed as follows:
|
wenzelm@17408
|
1577 |
|
wenzelm@17408
|
1578 |
assoc = uncurry (AList.lookup (op =))
|
wenzelm@17408
|
1579 |
assocs = these oo AList.lookup (op =)
|
wenzelm@17408
|
1580 |
overwrite = uncurry (AList.update (op =)) o swap
|
wenzelm@17408
|
1581 |
|
haftmann@17564
|
1582 |
* Pure/General: structure AList (cf. Pure/General/alist.ML) provides
|
haftmann@17564
|
1583 |
|
haftmann@17564
|
1584 |
val make: ('a -> 'b) -> 'a list -> ('a * 'b) list
|
haftmann@17564
|
1585 |
val find: ('a * 'b -> bool) -> ('c * 'b) list -> 'a -> 'c list
|
haftmann@17564
|
1586 |
|
haftmann@17564
|
1587 |
replacing make_keylist and keyfilter (occassionally used)
|
haftmann@17564
|
1588 |
Naive rewrites:
|
haftmann@17564
|
1589 |
|
haftmann@17564
|
1590 |
make_keylist = AList.make
|
haftmann@17564
|
1591 |
keyfilter = AList.find (op =)
|
haftmann@17564
|
1592 |
|
haftmann@17564
|
1593 |
* eq_fst and eq_snd now take explicit equality parameter, thus
|
haftmann@17564
|
1594 |
avoiding eqtypes. Naive rewrites:
|
haftmann@17564
|
1595 |
|
haftmann@17564
|
1596 |
eq_fst = eq_fst (op =)
|
haftmann@17564
|
1597 |
eq_snd = eq_snd (op =)
|
haftmann@17564
|
1598 |
|
haftmann@17564
|
1599 |
* Removed deprecated apl and apr (rarely used).
|
haftmann@17564
|
1600 |
Naive rewrites:
|
haftmann@17564
|
1601 |
|
haftmann@17564
|
1602 |
apl (n, op) =>>= curry op n
|
haftmann@17564
|
1603 |
apr (op, m) =>>= fn n => op (n, m)
|
haftmann@17564
|
1604 |
|
wenzelm@17408
|
1605 |
* Pure/General: structure OrdList (cf. Pure/General/ord_list.ML)
|
wenzelm@17408
|
1606 |
provides a reasonably efficient light-weight implementation of sets as
|
wenzelm@17408
|
1607 |
lists.
|
wenzelm@17408
|
1608 |
|
wenzelm@17408
|
1609 |
* Pure/General: generic tables (cf. Pure/General/table.ML) provide a
|
wenzelm@17408
|
1610 |
few new operations; existing lookup and update are now curried to
|
wenzelm@17408
|
1611 |
follow natural argument order (for use with fold etc.);
|
wenzelm@17408
|
1612 |
INCOMPATIBILITY, use (uncurry Symtab.lookup) etc. as last resort.
|
wenzelm@17408
|
1613 |
|
wenzelm@17408
|
1614 |
* Pure/General: output via the Isabelle channels of
|
wenzelm@17408
|
1615 |
writeln/warning/error etc. is now passed through Output.output, with a
|
wenzelm@17408
|
1616 |
hook for arbitrary transformations depending on the print_mode
|
wenzelm@17408
|
1617 |
(cf. Output.add_mode -- the first active mode that provides a output
|
wenzelm@17408
|
1618 |
function wins). Already formatted output may be embedded into further
|
wenzelm@17408
|
1619 |
text via Output.raw; the result of Pretty.string_of/str_of and derived
|
wenzelm@17408
|
1620 |
functions (string_of_term/cterm/thm etc.) is already marked raw to
|
wenzelm@17408
|
1621 |
accommodate easy composition of diagnostic messages etc. Programmers
|
wenzelm@17408
|
1622 |
rarely need to care about Output.output or Output.raw at all, with
|
wenzelm@17408
|
1623 |
some notable exceptions: Output.output is required when bypassing the
|
wenzelm@17408
|
1624 |
standard channels (writeln etc.), or in token translations to produce
|
wenzelm@17408
|
1625 |
properly formatted results; Output.raw is required when capturing
|
wenzelm@17408
|
1626 |
already output material that will eventually be presented to the user
|
wenzelm@17408
|
1627 |
a second time. For the default print mode, both Output.output and
|
wenzelm@17408
|
1628 |
Output.raw have no effect.
|
wenzelm@17408
|
1629 |
|
wenzelm@17408
|
1630 |
* Pure/General: Output.time_accumulator NAME creates an operator ('a
|
wenzelm@17408
|
1631 |
-> 'b) -> 'a -> 'b to measure runtime and count invocations; the
|
wenzelm@17408
|
1632 |
cumulative results are displayed at the end of a batch session.
|
wenzelm@17408
|
1633 |
|
wenzelm@17408
|
1634 |
* Pure/General: File.sysify_path and File.quote_sysify path have been
|
wenzelm@17408
|
1635 |
replaced by File.platform_path and File.shell_path (with appropriate
|
wenzelm@17408
|
1636 |
hooks). This provides a clean interface for unusual systems where the
|
wenzelm@17408
|
1637 |
internal and external process view of file names are different.
|
wenzelm@17408
|
1638 |
|
wenzelm@16689
|
1639 |
* Pure: more efficient orders for basic syntactic entities: added
|
wenzelm@16689
|
1640 |
fast_string_ord, fast_indexname_ord, fast_term_ord; changed sort_ord
|
wenzelm@16689
|
1641 |
and typ_ord to use fast_string_ord and fast_indexname_ord (term_ord is
|
wenzelm@16689
|
1642 |
NOT affected); structures Symtab, Vartab, Typtab, Termtab use the fast
|
wenzelm@16689
|
1643 |
orders now -- potential INCOMPATIBILITY for code that depends on a
|
wenzelm@16689
|
1644 |
particular order for Symtab.keys, Symtab.dest, etc. (consider using
|
wenzelm@16689
|
1645 |
Library.sort_strings on result).
|
wenzelm@16689
|
1646 |
|
wenzelm@17408
|
1647 |
* Pure/term.ML: combinators fold_atyps, fold_aterms, fold_term_types,
|
wenzelm@17408
|
1648 |
fold_types traverse types/terms from left to right, observing natural
|
wenzelm@17408
|
1649 |
argument order. Supercedes previous foldl_XXX versions, add_frees,
|
wenzelm@17408
|
1650 |
add_vars etc. have been adapted as well: INCOMPATIBILITY.
|
wenzelm@17408
|
1651 |
|
wenzelm@16151
|
1652 |
* Pure: name spaces have been refined, with significant changes of the
|
wenzelm@16234
|
1653 |
internal interfaces -- INCOMPATIBILITY. Renamed cond_extern(_table)
|
wenzelm@16234
|
1654 |
to extern(_table). The plain name entry path is superceded by a
|
wenzelm@16234
|
1655 |
general 'naming' context, which also includes the 'policy' to produce
|
wenzelm@16234
|
1656 |
a fully qualified name and external accesses of a fully qualified
|
wenzelm@16234
|
1657 |
name; NameSpace.extend is superceded by context dependent
|
wenzelm@16234
|
1658 |
Sign.declare_name. Several theory and proof context operations modify
|
wenzelm@16234
|
1659 |
the naming context. Especially note Theory.restore_naming and
|
wenzelm@16234
|
1660 |
ProofContext.restore_naming to get back to a sane state; note that
|
wenzelm@16234
|
1661 |
Theory.add_path is no longer sufficient to recover from
|
wenzelm@16234
|
1662 |
Theory.absolute_path in particular.
|
wenzelm@16234
|
1663 |
|
wenzelm@16234
|
1664 |
* Pure: new flags short_names (default false) and unique_names
|
wenzelm@16234
|
1665 |
(default true) for controlling output of qualified names. If
|
wenzelm@16234
|
1666 |
short_names is set, names are printed unqualified. If unique_names is
|
wenzelm@16234
|
1667 |
reset, the name prefix is reduced to the minimum required to achieve
|
wenzelm@16234
|
1668 |
the original result when interning again, even if there is an overlap
|
wenzelm@16234
|
1669 |
with earlier declarations.
|
wenzelm@16151
|
1670 |
|
wenzelm@16456
|
1671 |
* Pure/TheoryDataFun: change of the argument structure; 'prep_ext' is
|
wenzelm@16456
|
1672 |
now 'extend', and 'merge' gets an additional Pretty.pp argument
|
wenzelm@16456
|
1673 |
(useful for printing error messages). INCOMPATIBILITY.
|
wenzelm@16456
|
1674 |
|
wenzelm@16456
|
1675 |
* Pure: major reorganization of the theory context. Type Sign.sg and
|
wenzelm@16456
|
1676 |
Theory.theory are now identified, referring to the universal
|
wenzelm@16456
|
1677 |
Context.theory (see Pure/context.ML). Actual signature and theory
|
wenzelm@16456
|
1678 |
content is managed as theory data. The old code and interfaces were
|
wenzelm@16456
|
1679 |
spread over many files and structures; the new arrangement introduces
|
wenzelm@16456
|
1680 |
considerable INCOMPATIBILITY to gain more clarity:
|
wenzelm@16456
|
1681 |
|
wenzelm@16456
|
1682 |
Context -- theory management operations (name, identity, inclusion,
|
wenzelm@16456
|
1683 |
parents, ancestors, merge, etc.), plus generic theory data;
|
wenzelm@16456
|
1684 |
|
wenzelm@16456
|
1685 |
Sign -- logical signature and syntax operations (declaring consts,
|
wenzelm@16456
|
1686 |
types, etc.), plus certify/read for common entities;
|
wenzelm@16456
|
1687 |
|
wenzelm@16456
|
1688 |
Theory -- logical theory operations (stating axioms, definitions,
|
wenzelm@16456
|
1689 |
oracles), plus a copy of logical signature operations (consts,
|
wenzelm@16456
|
1690 |
types, etc.); also a few basic management operations (Theory.copy,
|
wenzelm@16456
|
1691 |
Theory.merge, etc.)
|
wenzelm@16456
|
1692 |
|
wenzelm@16456
|
1693 |
The most basic sign_of operations (Theory.sign_of, Thm.sign_of_thm
|
wenzelm@16456
|
1694 |
etc.) as well as the sign field in Thm.rep_thm etc. have been retained
|
wenzelm@16456
|
1695 |
for convenience -- they merely return the theory.
|
wenzelm@16456
|
1696 |
|
wenzelm@17193
|
1697 |
* Pure: type Type.tsig is superceded by theory in most interfaces.
|
wenzelm@17193
|
1698 |
|
wenzelm@16547
|
1699 |
* Pure: the Isar proof context type is already defined early in Pure
|
wenzelm@16547
|
1700 |
as Context.proof (note that ProofContext.context and Proof.context are
|
wenzelm@16547
|
1701 |
aliases, where the latter is the preferred name). This enables other
|
wenzelm@16547
|
1702 |
Isabelle components to refer to that type even before Isar is present.
|
wenzelm@16547
|
1703 |
|
wenzelm@16373
|
1704 |
* Pure/sign/theory: discontinued named name spaces (i.e. classK,
|
wenzelm@16373
|
1705 |
typeK, constK, axiomK, oracleK), but provide explicit operations for
|
wenzelm@16373
|
1706 |
any of these kinds. For example, Sign.intern typeK is now
|
wenzelm@16373
|
1707 |
Sign.intern_type, Theory.hide_space Sign.typeK is now
|
wenzelm@16373
|
1708 |
Theory.hide_types. Also note that former
|
wenzelm@16373
|
1709 |
Theory.hide_classes/types/consts are now
|
wenzelm@16373
|
1710 |
Theory.hide_classes_i/types_i/consts_i, while the non '_i' versions
|
wenzelm@16373
|
1711 |
internalize their arguments! INCOMPATIBILITY.
|
wenzelm@16373
|
1712 |
|
wenzelm@16506
|
1713 |
* Pure: get_thm interface (of PureThy and ProofContext) expects
|
wenzelm@16506
|
1714 |
datatype thmref (with constructors Name and NameSelection) instead of
|
wenzelm@16506
|
1715 |
plain string -- INCOMPATIBILITY;
|
wenzelm@16506
|
1716 |
|
wenzelm@16151
|
1717 |
* Pure: cases produced by proof methods specify options, where NONE
|
wenzelm@16234
|
1718 |
means to remove case bindings -- INCOMPATIBILITY in
|
wenzelm@16234
|
1719 |
(RAW_)METHOD_CASES.
|
wenzelm@16151
|
1720 |
|
wenzelm@16373
|
1721 |
* Pure: the following operations retrieve axioms or theorems from a
|
wenzelm@16373
|
1722 |
theory node or theory hierarchy, respectively:
|
wenzelm@16373
|
1723 |
|
wenzelm@16373
|
1724 |
Theory.axioms_of: theory -> (string * term) list
|
wenzelm@16373
|
1725 |
Theory.all_axioms_of: theory -> (string * term) list
|
wenzelm@16373
|
1726 |
PureThy.thms_of: theory -> (string * thm) list
|
wenzelm@16373
|
1727 |
PureThy.all_thms_of: theory -> (string * thm) list
|
wenzelm@16373
|
1728 |
|
wenzelm@16718
|
1729 |
* Pure: print_tac now outputs the goal through the trace channel.
|
wenzelm@16718
|
1730 |
|
wenzelm@17408
|
1731 |
* Isar toplevel: improved diagnostics, mostly for Poly/ML only.
|
wenzelm@17408
|
1732 |
Reference Toplevel.debug (default false) controls detailed printing
|
wenzelm@17408
|
1733 |
and tracing of low-level exceptions; Toplevel.profiling (default 0)
|
wenzelm@17408
|
1734 |
controls execution profiling -- set to 1 for time and 2 for space
|
wenzelm@17408
|
1735 |
(both increase the runtime).
|
wenzelm@17408
|
1736 |
|
wenzelm@17408
|
1737 |
* Isar session: The initial use of ROOT.ML is now always timed,
|
wenzelm@17408
|
1738 |
i.e. the log will show the actual process times, in contrast to the
|
wenzelm@17408
|
1739 |
elapsed wall-clock time that the outer shell wrapper produces.
|
wenzelm@17408
|
1740 |
|
wenzelm@17408
|
1741 |
* Simplifier: improved handling of bound variables (nameless
|
wenzelm@16997
|
1742 |
representation, avoid allocating new strings). Simprocs that invoke
|
wenzelm@16997
|
1743 |
the Simplifier recursively should use Simplifier.inherit_bounds to
|
wenzelm@17720
|
1744 |
avoid local name clashes. Failure to do so produces warnings
|
wenzelm@17720
|
1745 |
"Simplifier: renamed bound variable ..."; set Simplifier.debug_bounds
|
wenzelm@17720
|
1746 |
for further details.
|
wenzelm@16234
|
1747 |
|
wenzelm@17166
|
1748 |
* ML functions legacy_bindings and use_legacy_bindings produce ML fact
|
wenzelm@17166
|
1749 |
bindings for all theorems stored within a given theory; this may help
|
wenzelm@17166
|
1750 |
in porting non-Isar theories to Isar ones, while keeping ML proof
|
wenzelm@17166
|
1751 |
scripts for the time being.
|
wenzelm@17166
|
1752 |
|
wenzelm@17457
|
1753 |
* ML operator HTML.with_charset specifies the charset begin used for
|
wenzelm@17457
|
1754 |
generated HTML files. For example:
|
wenzelm@17457
|
1755 |
|
wenzelm@17457
|
1756 |
HTML.with_charset "utf-8" use_thy "Hebrew";
|
wenzelm@17538
|
1757 |
HTML.with_charset "utf-8" use_thy "Chinese";
|
wenzelm@17457
|
1758 |
|
wenzelm@16234
|
1759 |
|
wenzelm@16234
|
1760 |
*** System ***
|
wenzelm@16234
|
1761 |
|
wenzelm@16234
|
1762 |
* Allow symlinks to all proper Isabelle executables (Isabelle,
|
wenzelm@16234
|
1763 |
isabelle, isatool etc.).
|
wenzelm@16234
|
1764 |
|
wenzelm@16234
|
1765 |
* ISABELLE_DOC_FORMAT setting specifies preferred document format (for
|
wenzelm@16234
|
1766 |
isatool doc, isatool mkdir, display_drafts etc.).
|
wenzelm@16234
|
1767 |
|
wenzelm@16234
|
1768 |
* isatool usedir: option -f allows specification of the ML file to be
|
wenzelm@16234
|
1769 |
used by Isabelle; default is ROOT.ML.
|
wenzelm@16234
|
1770 |
|
wenzelm@16251
|
1771 |
* New isatool version outputs the version identifier of the Isabelle
|
wenzelm@16251
|
1772 |
distribution being used.
|
wenzelm@16251
|
1773 |
|
wenzelm@16251
|
1774 |
* HOL: new isatool dimacs2hol converts files in DIMACS CNF format
|
wenzelm@16234
|
1775 |
(containing Boolean satisfiability problems) into Isabelle/HOL
|
wenzelm@16234
|
1776 |
theories.
|
wenzelm@15703
|
1777 |
|
wenzelm@15703
|
1778 |
|
wenzelm@14655
|
1779 |
|
wenzelm@14606
|
1780 |
New in Isabelle2004 (April 2004)
|
wenzelm@14606
|
1781 |
--------------------------------
|
wenzelm@13280
|
1782 |
|
skalberg@14171
|
1783 |
*** General ***
|
skalberg@14171
|
1784 |
|
ballarin@14398
|
1785 |
* Provers/order.ML: new efficient reasoner for partial and linear orders.
|
ballarin@14398
|
1786 |
Replaces linorder.ML.
|
ballarin@14398
|
1787 |
|
wenzelm@14606
|
1788 |
* Pure: Greek letters (except small lambda, \<lambda>), as well as Gothic
|
wenzelm@14606
|
1789 |
(\<aa>...\<zz>\<AA>...\<ZZ>), calligraphic (\<A>...\<Z>), and Euler
|
skalberg@14173
|
1790 |
(\<a>...\<z>), are now considered normal letters, and can therefore
|
skalberg@14173
|
1791 |
be used anywhere where an ASCII letter (a...zA...Z) has until
|
skalberg@14173
|
1792 |
now. COMPATIBILITY: This obviously changes the parsing of some
|
skalberg@14173
|
1793 |
terms, especially where a symbol has been used as a binder, say
|
skalberg@14173
|
1794 |
'\<Pi>x. ...', which is now a type error since \<Pi>x will be parsed
|
skalberg@14173
|
1795 |
as an identifier. Fix it by inserting a space around former
|
skalberg@14173
|
1796 |
symbols. Call 'isatool fixgreek' to try to fix parsing errors in
|
skalberg@14173
|
1797 |
existing theory and ML files.
|
skalberg@14171
|
1798 |
|
paulson@14237
|
1799 |
* Pure: Macintosh and Windows line-breaks are now allowed in theory files.
|
paulson@14237
|
1800 |
|
wenzelm@14731
|
1801 |
* Pure: single letter sub/superscripts (\<^isub> and \<^isup>) are now
|
wenzelm@14731
|
1802 |
allowed in identifiers. Similar to Greek letters \<^isub> is now considered
|
wenzelm@14731
|
1803 |
a normal (but invisible) letter. For multiple letter subscripts repeat
|
wenzelm@14731
|
1804 |
\<^isub> like this: x\<^isub>1\<^isub>2.
|
kleing@14233
|
1805 |
|
kleing@14333
|
1806 |
* Pure: There are now sub-/superscripts that can span more than one
|
kleing@14333
|
1807 |
character. Text between \<^bsub> and \<^esub> is set in subscript in
|
wenzelm@14606
|
1808 |
ProofGeneral and LaTeX, text between \<^bsup> and \<^esup> in
|
wenzelm@14606
|
1809 |
superscript. The new control characters are not identifier parts.
|
kleing@14333
|
1810 |
|
schirmer@14561
|
1811 |
* Pure: Control-symbols of the form \<^raw:...> will literally print the
|
wenzelm@14606
|
1812 |
content of "..." to the latex file instead of \isacntrl... . The "..."
|
wenzelm@14606
|
1813 |
may consist of any printable characters excluding the end bracket >.
|
schirmer@14361
|
1814 |
|
paulson@14237
|
1815 |
* Pure: Using new Isar command "finalconsts" (or the ML functions
|
paulson@14237
|
1816 |
Theory.add_finals or Theory.add_finals_i) it is now possible to
|
paulson@14237
|
1817 |
declare constants "final", which prevents their being given a definition
|
paulson@14237
|
1818 |
later. It is useful for constants whose behaviour is fixed axiomatically
|
skalberg@14224
|
1819 |
rather than definitionally, such as the meta-logic connectives.
|
skalberg@14224
|
1820 |
|
wenzelm@14606
|
1821 |
* Pure: 'instance' now handles general arities with general sorts
|
wenzelm@14606
|
1822 |
(i.e. intersections of classes),
|
skalberg@14503
|
1823 |
|
kleing@14547
|
1824 |
* Presentation: generated HTML now uses a CSS style sheet to make layout
|
wenzelm@14731
|
1825 |
(somewhat) independent of content. It is copied from lib/html/isabelle.css.
|
kleing@14547
|
1826 |
It can be changed to alter the colors/layout of generated pages.
|
kleing@14547
|
1827 |
|
wenzelm@14556
|
1828 |
|
ballarin@14175
|
1829 |
*** Isar ***
|
ballarin@14175
|
1830 |
|
ballarin@14508
|
1831 |
* Tactic emulation methods rule_tac, erule_tac, drule_tac, frule_tac,
|
ballarin@14508
|
1832 |
cut_tac, subgoal_tac and thin_tac:
|
ballarin@14175
|
1833 |
- Now understand static (Isar) contexts. As a consequence, users of Isar
|
ballarin@14175
|
1834 |
locales are no longer forced to write Isar proof scripts.
|
ballarin@14175
|
1835 |
For details see Isar Reference Manual, paragraph 4.3.2: Further tactic
|
ballarin@14175
|
1836 |
emulations.
|
ballarin@14175
|
1837 |
- INCOMPATIBILITY: names of variables to be instantiated may no
|
ballarin@14211
|
1838 |
longer be enclosed in quotes. Instead, precede variable name with `?'.
|
ballarin@14211
|
1839 |
This is consistent with the instantiation attribute "where".
|
ballarin@14211
|
1840 |
|
ballarin@14257
|
1841 |
* Attributes "where" and "of":
|
ballarin@14285
|
1842 |
- Now take type variables of instantiated theorem into account when reading
|
ballarin@14285
|
1843 |
the instantiation string. This fixes a bug that caused instantiated
|
ballarin@14285
|
1844 |
theorems to have too special types in some circumstances.
|
ballarin@14285
|
1845 |
- "where" permits explicit instantiations of type variables.
|
ballarin@14257
|
1846 |
|
wenzelm@14556
|
1847 |
* Calculation commands "moreover" and "also" no longer interfere with
|
wenzelm@14556
|
1848 |
current facts ("this"), admitting arbitrary combinations with "then"
|
wenzelm@14556
|
1849 |
and derived forms.
|
kleing@14283
|
1850 |
|
ballarin@14211
|
1851 |
* Locales:
|
ballarin@14211
|
1852 |
- Goal statements involving the context element "includes" no longer
|
ballarin@14211
|
1853 |
generate theorems with internal delta predicates (those ending on
|
ballarin@14211
|
1854 |
"_axioms") in the premise.
|
ballarin@14211
|
1855 |
Resolve particular premise with <locale>.intro to obtain old form.
|
ballarin@14211
|
1856 |
- Fixed bug in type inference ("unify_frozen") that prevented mix of target
|
ballarin@14211
|
1857 |
specification and "includes" elements in goal statement.
|
ballarin@14254
|
1858 |
- Rule sets <locale>.intro and <locale>.axioms no longer declared as
|
ballarin@14254
|
1859 |
[intro?] and [elim?] (respectively) by default.
|
ballarin@14508
|
1860 |
- Experimental command for instantiation of locales in proof contexts:
|
ballarin@14551
|
1861 |
instantiate <label>[<attrs>]: <loc>
|
ballarin@14508
|
1862 |
Instantiates locale <loc> and adds all its theorems to the current context
|
ballarin@14551
|
1863 |
taking into account their attributes. Label and attrs are optional
|
ballarin@14551
|
1864 |
modifiers, like in theorem declarations. If present, names of
|
ballarin@14551
|
1865 |
instantiated theorems are qualified with <label>, and the attributes
|
ballarin@14551
|
1866 |
<attrs> are applied after any attributes these theorems might have already.
|
ballarin@14551
|
1867 |
If the locale has assumptions, a chained fact of the form
|
ballarin@14508
|
1868 |
"<loc> t1 ... tn" is expected from which instantiations of the parameters
|
ballarin@14551
|
1869 |
are derived. The command does not support old-style locales declared
|
ballarin@14551
|
1870 |
with "locale (open)".
|
ballarin@14551
|
1871 |
A few (very simple) examples can be found in FOL/ex/LocaleInst.thy.
|
ballarin@14175
|
1872 |
|
ballarin@14175
|
1873 |
* HOL: Tactic emulation methods induct_tac and case_tac understand static
|
ballarin@14175
|
1874 |
(Isar) contexts.
|
ballarin@14175
|
1875 |
|
wenzelm@14556
|
1876 |
|
kleing@14136
|
1877 |
*** HOL ***
|
kleing@14136
|
1878 |
|
kleing@14624
|
1879 |
* Proof import: new image HOL4 contains the imported library from
|
kleing@14624
|
1880 |
the HOL4 system with about 2500 theorems. It is imported by
|
kleing@14624
|
1881 |
replaying proof terms produced by HOL4 in Isabelle. The HOL4 image
|
kleing@14624
|
1882 |
can be used like any other Isabelle image. See
|
kleing@14624
|
1883 |
HOL/Import/HOL/README for more information.
|
kleing@14624
|
1884 |
|
ballarin@14398
|
1885 |
* Simplifier:
|
ballarin@14398
|
1886 |
- Much improved handling of linear and partial orders.
|
ballarin@14398
|
1887 |
Reasoners for linear and partial orders are set up for type classes
|
ballarin@14398
|
1888 |
"linorder" and "order" respectively, and are added to the default simpset
|
ballarin@14398
|
1889 |
as solvers. This means that the simplifier can build transitivity chains
|
ballarin@14398
|
1890 |
to solve goals from the assumptions.
|
ballarin@14398
|
1891 |
- INCOMPATIBILITY: old proofs break occasionally. Typically, applications
|
ballarin@14398
|
1892 |
of blast or auto after simplification become unnecessary because the goal
|
ballarin@14398
|
1893 |
is solved by simplification already.
|
ballarin@14398
|
1894 |
|
wenzelm@14731
|
1895 |
* Numerics: new theory Ring_and_Field contains over 250 basic numerical laws,
|
paulson@14389
|
1896 |
all proved in axiomatic type classes for semirings, rings and fields.
|
paulson@14389
|
1897 |
|
paulson@14389
|
1898 |
* Numerics:
|
paulson@14389
|
1899 |
- Numeric types (nat, int, and in HOL-Complex rat, real, complex, etc.) are
|
wenzelm@14731
|
1900 |
now formalized using the Ring_and_Field theory mentioned above.
|
paulson@14389
|
1901 |
- INCOMPATIBILITY: simplification and arithmetic behaves somewhat differently
|
paulson@14389
|
1902 |
than before, because now they are set up once in a generic manner.
|
wenzelm@14731
|
1903 |
- INCOMPATIBILITY: many type-specific arithmetic laws have gone.
|
paulson@14480
|
1904 |
Look for the general versions in Ring_and_Field (and Power if they concern
|
paulson@14480
|
1905 |
exponentiation).
|
paulson@14389
|
1906 |
|
paulson@14401
|
1907 |
* Type "rat" of the rational numbers is now available in HOL-Complex.
|
paulson@14389
|
1908 |
|
schirmer@14255
|
1909 |
* Records:
|
schirmer@14255
|
1910 |
- Record types are now by default printed with their type abbreviation
|
schirmer@14255
|
1911 |
instead of the list of all field types. This can be configured via
|
schirmer@14255
|
1912 |
the reference "print_record_type_abbr".
|
wenzelm@14731
|
1913 |
- Simproc "record_upd_simproc" for simplification of multiple updates added
|
schirmer@14255
|
1914 |
(not enabled by default).
|
schirmer@14427
|
1915 |
- Simproc "record_ex_sel_eq_simproc" to simplify EX x. sel r = x resp.
|
schirmer@14427
|
1916 |
EX x. x = sel r to True (not enabled by default).
|
schirmer@14255
|
1917 |
- Tactic "record_split_simp_tac" to split and simplify records added.
|
wenzelm@14731
|
1918 |
|
kleing@14136
|
1919 |
* 'specification' command added, allowing for definition by
|
skalberg@14224
|
1920 |
specification. There is also an 'ax_specification' command that
|
skalberg@14224
|
1921 |
introduces the new constants axiomatically.
|
kleing@14136
|
1922 |
|
nipkow@14375
|
1923 |
* arith(_tac) is now able to generate counterexamples for reals as well.
|
nipkow@14375
|
1924 |
|
ballarin@14399
|
1925 |
* HOL-Algebra: new locale "ring" for non-commutative rings.
|
ballarin@14399
|
1926 |
|
paulson@14243
|
1927 |
* HOL-ex: InductiveInvariant_examples illustrates advanced recursive function
|
kleing@14610
|
1928 |
definitions, thanks to Sava Krsti\'{c} and John Matthews.
|
kleing@14610
|
1929 |
|
wenzelm@14731
|
1930 |
* HOL-Matrix: a first theory for matrices in HOL with an application of
|
kleing@14610
|
1931 |
matrix theory to linear programming.
|
kleing@14136
|
1932 |
|
nipkow@14380
|
1933 |
* Unions and Intersections:
|
nipkow@15119
|
1934 |
The latex output syntax of UN and INT has been changed
|
nipkow@15119
|
1935 |
from "\Union x \in A. B" to "\Union_{x \in A} B"
|
nipkow@15119
|
1936 |
i.e. the index formulae has become a subscript.
|
nipkow@15119
|
1937 |
Similarly for "\Union x. B", and for \Inter instead of \Union.
|
nipkow@14380
|
1938 |
|
kleing@14418
|
1939 |
* Unions and Intersections over Intervals:
|
wenzelm@14731
|
1940 |
There is new short syntax "UN i<=n. A" for "UN i:{0..n}. A". There is
|
wenzelm@14731
|
1941 |
also an x-symbol version with subscripts "\<Union>\<^bsub>i <= n\<^esub>. A"
|
kleing@14418
|
1942 |
like in normal math, and corresponding versions for < and for intersection.
|
kleing@14418
|
1943 |
|
nipkow@15677
|
1944 |
* HOL/List: Ordering "lexico" is renamed "lenlex" and the standard
|
nipkow@15677
|
1945 |
lexicographic dictonary ordering has been added as "lexord".
|
nipkow@15677
|
1946 |
|
paulson@14401
|
1947 |
* ML: the legacy theory structures Int and List have been removed. They had
|
paulson@14401
|
1948 |
conflicted with ML Basis Library structures having the same names.
|
nipkow@14380
|
1949 |
|
webertj@14464
|
1950 |
* 'refute' command added to search for (finite) countermodels. Only works
|
webertj@14464
|
1951 |
for a fragment of HOL. The installation of an external SAT solver is
|
webertj@14464
|
1952 |
highly recommended. See "HOL/Refute.thy" for details.
|
webertj@14464
|
1953 |
|
berghofe@14602
|
1954 |
* 'quickcheck' command: Allows to find counterexamples by evaluating
|
berghofe@14602
|
1955 |
formulae under an assignment of free variables to random values.
|
berghofe@14602
|
1956 |
In contrast to 'refute', it can deal with inductive datatypes,
|
berghofe@14602
|
1957 |
but cannot handle quantifiers. See "HOL/ex/Quickcheck_Examples.thy"
|
berghofe@14602
|
1958 |
for examples.
|
webertj@14464
|
1959 |
|
wenzelm@14606
|
1960 |
|
oheimb@14536
|
1961 |
*** HOLCF ***
|
oheimb@14536
|
1962 |
|
oheimb@14536
|
1963 |
* Streams now come with concatenation and are part of the HOLCF image
|
oheimb@14536
|
1964 |
|
wenzelm@14572
|
1965 |
|
wenzelm@14572
|
1966 |
|
kleing@14136
|
1967 |
New in Isabelle2003 (May 2003)
|
wenzelm@14606
|
1968 |
------------------------------
|
kleing@14136
|
1969 |
|
wenzelm@13280
|
1970 |
*** General ***
|
wenzelm@13280
|
1971 |
|
berghofe@13618
|
1972 |
* Provers/simplifier:
|
berghofe@13618
|
1973 |
|
nipkow@13781
|
1974 |
- Completely reimplemented method simp (ML: Asm_full_simp_tac):
|
berghofe@13618
|
1975 |
Assumptions are now subject to complete mutual simplification,
|
berghofe@13618
|
1976 |
not just from left to right. The simplifier now preserves
|
berghofe@13618
|
1977 |
the order of assumptions.
|
berghofe@13618
|
1978 |
|
berghofe@13618
|
1979 |
Potential INCOMPATIBILITY:
|
berghofe@13618
|
1980 |
|
nipkow@13781
|
1981 |
-- simp sometimes diverges where the old version did
|
nipkow@13781
|
1982 |
not, e.g. invoking simp on the goal
|
berghofe@13618
|
1983 |
|
berghofe@13618
|
1984 |
[| P (f x); y = x; f x = f y |] ==> Q
|
berghofe@13618
|
1985 |
|
nipkow@13781
|
1986 |
now gives rise to the infinite reduction sequence
|
nipkow@13781
|
1987 |
|
nipkow@13781
|
1988 |
P(f x) --(f x = f y)--> P(f y) --(y = x)--> P(f x) --(f x = f y)--> ...
|
nipkow@13781
|
1989 |
|
nipkow@13781
|
1990 |
Using "simp (asm_lr)" (ML: Asm_lr_simp_tac) instead often solves this
|
nipkow@13781
|
1991 |
kind of problem.
|
nipkow@13781
|
1992 |
|
nipkow@13781
|
1993 |
-- Tactics combining classical reasoner and simplification (such as auto)
|
nipkow@13781
|
1994 |
are also affected by this change, because many of them rely on
|
nipkow@13781
|
1995 |
simp. They may sometimes diverge as well or yield a different numbers
|
nipkow@13781
|
1996 |
of subgoals. Try to use e.g. force, fastsimp, or safe instead of auto
|
nipkow@13781
|
1997 |
in case of problems. Sometimes subsequent calls to the classical
|
nipkow@13781
|
1998 |
reasoner will fail because a preceeding call to the simplifier too
|
nipkow@13781
|
1999 |
eagerly simplified the goal, e.g. deleted redundant premises.
|
berghofe@13618
|
2000 |
|
berghofe@13618
|
2001 |
- The simplifier trace now shows the names of the applied rewrite rules
|
berghofe@13618
|
2002 |
|
nipkow@13829
|
2003 |
- You can limit the number of recursive invocations of the simplifier
|
nipkow@13829
|
2004 |
during conditional rewriting (where the simplifie tries to solve the
|
nipkow@13829
|
2005 |
conditions before applying the rewrite rule):
|
nipkow@13829
|
2006 |
ML "simp_depth_limit := n"
|
nipkow@13829
|
2007 |
where n is an integer. Thus you can force termination where previously
|
nipkow@13829
|
2008 |
the simplifier would diverge.
|
nipkow@13829
|
2009 |
|
ballarin@13835
|
2010 |
- Accepts free variables as head terms in congruence rules. Useful in Isar.
|
nipkow@13829
|
2011 |
|
ballarin@13938
|
2012 |
- No longer aborts on failed congruence proof. Instead, the
|
ballarin@13938
|
2013 |
congruence is ignored.
|
ballarin@13938
|
2014 |
|
berghofe@14008
|
2015 |
* Pure: New generic framework for extracting programs from constructive
|
berghofe@14008
|
2016 |
proofs. See HOL/Extraction.thy for an example instantiation, as well
|
berghofe@14008
|
2017 |
as HOL/Extraction for some case studies.
|
berghofe@14008
|
2018 |
|
nipkow@13868
|
2019 |
* Pure: The main goal of the proof state is no longer shown by default, only
|
nipkow@13868
|
2020 |
the subgoals. This behaviour is controlled by a new flag.
|
ballarin@13835
|
2021 |
PG menu: Isabelle/Isar -> Settings -> Show Main Goal
|
nipkow@13815
|
2022 |
(ML: Proof.show_main_goal).
|
nipkow@13815
|
2023 |
|
nipkow@13815
|
2024 |
* Pure: You can find all matching introduction rules for subgoal 1, i.e. all
|
nipkow@13815
|
2025 |
rules whose conclusion matches subgoal 1:
|
nipkow@13815
|
2026 |
PG menu: Isabelle/Isar -> Show me -> matching rules
|
nipkow@13815
|
2027 |
The rules are ordered by how closely they match the subgoal.
|
nipkow@13815
|
2028 |
In particular, rules that solve a subgoal outright are displayed first
|
nipkow@13815
|
2029 |
(or rather last, the way they are printed).
|
nipkow@13815
|
2030 |
(ML: ProofGeneral.print_intros())
|
nipkow@13815
|
2031 |
|
nipkow@13815
|
2032 |
* Pure: New flag trace_unify_fail causes unification to print
|
nipkow@13781
|
2033 |
diagnostic information (PG: in trace buffer) when it fails. This is
|
nipkow@13781
|
2034 |
useful for figuring out why single step proofs like rule, erule or
|
nipkow@13781
|
2035 |
assumption failed.
|
nipkow@13781
|
2036 |
|
nipkow@13815
|
2037 |
* Pure: Locale specifications now produce predicate definitions
|
wenzelm@13410
|
2038 |
according to the body of text (covering assumptions modulo local
|
wenzelm@13410
|
2039 |
definitions); predicate "loc_axioms" covers newly introduced text,
|
wenzelm@13410
|
2040 |
while "loc" is cumulative wrt. all included locale expressions; the
|
wenzelm@13410
|
2041 |
latter view is presented only on export into the global theory
|
wenzelm@13410
|
2042 |
context; potential INCOMPATIBILITY, use "(open)" option to fall back
|
wenzelm@13410
|
2043 |
on the old view without predicates;
|
wenzelm@13410
|
2044 |
|
wenzelm@13459
|
2045 |
* Pure: predefined locales "var" and "struct" are useful for sharing
|
wenzelm@13459
|
2046 |
parameters (as in CASL, for example); just specify something like
|
wenzelm@13459
|
2047 |
``var x + var y + struct M'' as import;
|
wenzelm@13459
|
2048 |
|
wenzelm@13463
|
2049 |
* Pure: improved thms_containing: proper indexing of facts instead of
|
wenzelm@13463
|
2050 |
raw theorems; check validity of results wrt. current name space;
|
wenzelm@13463
|
2051 |
include local facts of proof configuration (also covers active
|
wenzelm@13541
|
2052 |
locales), cover fixed variables in index; may use "_" in term
|
wenzelm@13541
|
2053 |
specification; an optional limit for the number of printed facts may
|
wenzelm@13541
|
2054 |
be given (the default is 40);
|
wenzelm@13541
|
2055 |
|
wenzelm@13541
|
2056 |
* Pure: disallow duplicate fact bindings within new-style theory files
|
wenzelm@13541
|
2057 |
(batch-mode only);
|
wenzelm@13540
|
2058 |
|
wenzelm@13463
|
2059 |
* Provers: improved induct method: assumptions introduced by case
|
wenzelm@13463
|
2060 |
"foo" are split into "foo.hyps" (from the rule) and "foo.prems" (from
|
wenzelm@13463
|
2061 |
the goal statement); "foo" still refers to all facts collectively;
|
wenzelm@13463
|
2062 |
|
paulson@13550
|
2063 |
* Provers: the function blast.overloaded has been removed: all constants
|
paulson@13550
|
2064 |
are regarded as potentially overloaded, which improves robustness in exchange
|
paulson@13550
|
2065 |
for slight decrease in efficiency;
|
paulson@13550
|
2066 |
|
nipkow@13781
|
2067 |
* Provers/linorder: New generic prover for transitivity reasoning over
|
nipkow@13781
|
2068 |
linear orders. Note: this prover is not efficient!
|
nipkow@13781
|
2069 |
|
wenzelm@13522
|
2070 |
* Isar: preview of problems to finish 'show' now produce an error
|
wenzelm@13522
|
2071 |
rather than just a warning (in interactive mode);
|
wenzelm@13522
|
2072 |
|
wenzelm@13280
|
2073 |
|
nipkow@13158
|
2074 |
*** HOL ***
|
nipkow@13158
|
2075 |
|
nipkow@13899
|
2076 |
* arith(_tac)
|
nipkow@13899
|
2077 |
|
nipkow@13899
|
2078 |
- Produces a counter example if it cannot prove a goal.
|
nipkow@13899
|
2079 |
Note that the counter example may be spurious if the goal is not a formula
|
nipkow@13899
|
2080 |
of quantifier-free linear arithmetic.
|
nipkow@13899
|
2081 |
In ProofGeneral the counter example appears in the trace buffer.
|
nipkow@13899
|
2082 |
|
nipkow@13899
|
2083 |
- Knows about div k and mod k where k is a numeral of type nat or int.
|
nipkow@13899
|
2084 |
|
nipkow@13899
|
2085 |
- Calls full Presburger arithmetic (by Amine Chaieb) if quantifier-free
|
nipkow@13899
|
2086 |
linear arithmetic fails. This takes account of quantifiers and divisibility.
|
wenzelm@14731
|
2087 |
Presburger arithmetic can also be called explicitly via presburger(_tac).
|
nipkow@13899
|
2088 |
|
nipkow@13899
|
2089 |
* simp's arithmetic capabilities have been enhanced a bit: it now
|
nipkow@13899
|
2090 |
takes ~= in premises into account (by performing a case split);
|
nipkow@13899
|
2091 |
|
nipkow@13899
|
2092 |
* simp reduces "m*(n div m) + n mod m" to n, even if the two summands
|
nipkow@13899
|
2093 |
are distributed over a sum of terms;
|
nipkow@13899
|
2094 |
|
ballarin@13735
|
2095 |
* New tactic "trans_tac" and method "trans" instantiate
|
ballarin@13735
|
2096 |
Provers/linorder.ML for axclasses "order" and "linorder" (predicates
|
wenzelm@14731
|
2097 |
"<=", "<" and "=").
|
wenzelm@14731
|
2098 |
|
wenzelm@14731
|
2099 |
* function INCOMPATIBILITIES: Pi-sets have been redefined and moved from main
|
paulson@13587
|
2100 |
HOL to Library/FuncSet; constant "Fun.op o" is now called "Fun.comp";
|
paulson@13587
|
2101 |
|
wenzelm@13443
|
2102 |
* 'typedef' command has new option "open" to suppress the set
|
wenzelm@13443
|
2103 |
definition;
|
wenzelm@13443
|
2104 |
|
wenzelm@13522
|
2105 |
* functions Min and Max on finite sets have been introduced (theory
|
wenzelm@13522
|
2106 |
Finite_Set);
|
nipkow@13492
|
2107 |
|
wenzelm@13443
|
2108 |
* attribute [symmetric] now works for relations as well; it turns
|
wenzelm@13443
|
2109 |
(x,y) : R^-1 into (y,x) : R, and vice versa;
|
wenzelm@13443
|
2110 |
|
nipkow@13613
|
2111 |
* induct over a !!-quantified statement (say !!x1..xn):
|
nipkow@13613
|
2112 |
each "case" automatically performs "fix x1 .. xn" with exactly those names.
|
nipkow@13613
|
2113 |
|
nipkow@13899
|
2114 |
* Map: `empty' is no longer a constant but a syntactic abbreviation for
|
nipkow@13899
|
2115 |
%x. None. Warning: empty_def now refers to the previously hidden definition
|
nipkow@13899
|
2116 |
of the empty set.
|
nipkow@13899
|
2117 |
|
ballarin@14018
|
2118 |
* Algebra: formalization of classical algebra. Intended as base for
|
ballarin@14018
|
2119 |
any algebraic development in Isabelle. Currently covers group theory
|
ballarin@14018
|
2120 |
(up to Sylow's theorem) and ring theory (Universal Property of
|
ballarin@14018
|
2121 |
Univariate Polynomials). Contributions welcome;
|
paulson@13960
|
2122 |
|
paulson@13960
|
2123 |
* GroupTheory: deleted, since its material has been moved to Algebra;
|
paulson@13960
|
2124 |
|
wenzelm@14731
|
2125 |
* Complex: new directory of the complex numbers with numeric constants,
|
wenzelm@14731
|
2126 |
nonstandard complex numbers, and some complex analysis, standard and
|
paulson@13966
|
2127 |
nonstandard (Jacques Fleuriot);
|
paulson@13966
|
2128 |
|
paulson@13966
|
2129 |
* HOL-Complex: new image for analysis, replacing HOL-Real and HOL-Hyperreal;
|
paulson@13966
|
2130 |
|
wenzelm@14731
|
2131 |
* Hyperreal: introduced Gauge integration and hyperreal logarithms (Jacques
|
paulson@13966
|
2132 |
Fleuriot);
|
paulson@13960
|
2133 |
|
wenzelm@13549
|
2134 |
* Real/HahnBanach: updated and adapted to locales;
|
wenzelm@13549
|
2135 |
|
ballarin@13995
|
2136 |
* NumberTheory: added Gauss's law of quadratic reciprocity (by Avigad,
|
ballarin@13995
|
2137 |
Gray and Kramer);
|
paulson@13872
|
2138 |
|
paulson@13872
|
2139 |
* UNITY: added the Meier-Sanders theory of progress sets;
|
paulson@13872
|
2140 |
|
kleing@14011
|
2141 |
* MicroJava: bytecode verifier and lightweight bytecode verifier
|
kleing@14011
|
2142 |
as abstract algorithms, instantiated to the JVM;
|
kleing@14011
|
2143 |
|
schirmer@14010
|
2144 |
* Bali: Java source language formalization. Type system, operational
|
schirmer@14010
|
2145 |
semantics, axiomatic semantics. Supported language features:
|
schirmer@14010
|
2146 |
classes, interfaces, objects,virtual methods, static methods,
|
schirmer@14010
|
2147 |
static/instance fields, arrays, access modifiers, definite
|
schirmer@14010
|
2148 |
assignment, exceptions.
|
wenzelm@13549
|
2149 |
|
kleing@14011
|
2150 |
|
wenzelm@13549
|
2151 |
*** ZF ***
|
wenzelm@13549
|
2152 |
|
webertj@15154
|
2153 |
* ZF/Constructible: consistency proof for AC (Gdel's constructible
|
wenzelm@13549
|
2154 |
universe, etc.);
|
wenzelm@13549
|
2155 |
|
paulson@13872
|
2156 |
* Main ZF: virtually all theories converted to new-style format;
|
nipkow@13518
|
2157 |
|
wenzelm@13280
|
2158 |
|
wenzelm@13478
|
2159 |
*** ML ***
|
wenzelm@13478
|
2160 |
|
wenzelm@13478
|
2161 |
* Pure: Tactic.prove provides sane interface for internal proofs;
|
wenzelm@13478
|
2162 |
omits the infamous "standard" operation, so this is more appropriate
|
wenzelm@13478
|
2163 |
than prove_goalw_cterm in many situations (e.g. in simprocs);
|
wenzelm@13478
|
2164 |
|
wenzelm@13478
|
2165 |
* Pure: improved error reporting of simprocs;
|
wenzelm@13478
|
2166 |
|
wenzelm@13478
|
2167 |
* Provers: Simplifier.simproc(_i) provides sane interface for setting
|
wenzelm@13478
|
2168 |
up simprocs;
|
wenzelm@13478
|
2169 |
|
wenzelm@13478
|
2170 |
|
kleing@13953
|
2171 |
*** Document preparation ***
|
kleing@13953
|
2172 |
|
kleing@13953
|
2173 |
* uses \par instead of \\ for line breaks in theory text. This may
|
kleing@13953
|
2174 |
shift some page breaks in large documents. To get the old behaviour
|
kleing@13953
|
2175 |
use \renewcommand{\isanewline}{\mbox{}\\\mbox{}} in root.tex.
|
kleing@13953
|
2176 |
|
wenzelm@14731
|
2177 |
* minimized dependencies of isabelle.sty and isabellesym.sty on
|
kleing@13953
|
2178 |
other packages
|
kleing@13953
|
2179 |
|
kleing@13953
|
2180 |
* \<euro> now needs package babel/greek instead of marvosym (which
|
kleing@13953
|
2181 |
broke \Rightarrow)
|
kleing@13953
|
2182 |
|
wenzelm@14731
|
2183 |
* normal size for \<zero>...\<nine> (uses \mathbf instead of
|
kleing@13954
|
2184 |
textcomp package)
|
kleing@13953
|
2185 |
|
wenzelm@13280
|
2186 |
|
wenzelm@14572
|
2187 |
|
wenzelm@12984
|
2188 |
New in Isabelle2002 (March 2002)
|
wenzelm@12984
|
2189 |
--------------------------------
|
wenzelm@11474
|
2190 |
|
wenzelm@11572
|
2191 |
*** Document preparation ***
|
wenzelm@11572
|
2192 |
|
wenzelm@11842
|
2193 |
* greatly simplified document preparation setup, including more
|
wenzelm@11842
|
2194 |
graceful interpretation of isatool usedir -i/-d/-D options, and more
|
wenzelm@11842
|
2195 |
instructive isatool mkdir; users should basically be able to get
|
wenzelm@12899
|
2196 |
started with "isatool mkdir HOL Test && isatool make"; alternatively,
|
wenzelm@12899
|
2197 |
users may run a separate document processing stage manually like this:
|
wenzelm@12899
|
2198 |
"isatool usedir -D output HOL Test && isatool document Test/output";
|
wenzelm@11842
|
2199 |
|
wenzelm@11842
|
2200 |
* theory dependency graph may now be incorporated into documents;
|
wenzelm@11842
|
2201 |
isatool usedir -g true will produce session_graph.eps/.pdf for use
|
wenzelm@11842
|
2202 |
with \includegraphics of LaTeX;
|
wenzelm@11842
|
2203 |
|
wenzelm@11864
|
2204 |
* proper spacing of consecutive markup elements, especially text
|
wenzelm@11864
|
2205 |
blocks after section headings;
|
wenzelm@11864
|
2206 |
|
wenzelm@11572
|
2207 |
* support bold style (for single symbols only), input syntax is like
|
wenzelm@11572
|
2208 |
this: "\<^bold>\<alpha>" or "\<^bold>A";
|
wenzelm@11572
|
2209 |
|
wenzelm@11814
|
2210 |
* \<bullet> is now output as bold \cdot by default, which looks much
|
wenzelm@11572
|
2211 |
better in printed text;
|
wenzelm@11572
|
2212 |
|
wenzelm@11712
|
2213 |
* added default LaTeX bindings for \<tturnstile> and \<TTurnstile>;
|
wenzelm@11712
|
2214 |
note that these symbols are currently unavailable in Proof General /
|
wenzelm@12769
|
2215 |
X-Symbol; new symbols \<zero>, \<one>, ..., \<nine>, and \<euro>;
|
wenzelm@12690
|
2216 |
|
wenzelm@12690
|
2217 |
* isatool latex no longer depends on changed TEXINPUTS, instead
|
wenzelm@12690
|
2218 |
isatool document copies the Isabelle style files to the target
|
wenzelm@12690
|
2219 |
location;
|
wenzelm@11712
|
2220 |
|
wenzelm@11572
|
2221 |
|
wenzelm@11633
|
2222 |
*** Isar ***
|
wenzelm@11633
|
2223 |
|
wenzelm@12312
|
2224 |
* Pure/Provers: improved proof by cases and induction;
|
wenzelm@12280
|
2225 |
- 'case' command admits impromptu naming of parameters (such as
|
wenzelm@12280
|
2226 |
"case (Suc n)");
|
wenzelm@12280
|
2227 |
- 'induct' method divinates rule instantiation from the inductive
|
wenzelm@12280
|
2228 |
claim; no longer requires excessive ?P bindings for proper
|
wenzelm@12280
|
2229 |
instantiation of cases;
|
wenzelm@12280
|
2230 |
- 'induct' method properly enumerates all possibilities of set/type
|
wenzelm@12280
|
2231 |
rules; as a consequence facts may be also passed through *type*
|
wenzelm@12280
|
2232 |
rules without further ado;
|
wenzelm@12280
|
2233 |
- 'induct' method now derives symbolic cases from the *rulified*
|
wenzelm@12280
|
2234 |
rule (before it used to rulify cases stemming from the internal
|
wenzelm@12280
|
2235 |
atomized version); this means that the context of a non-atomic
|
wenzelm@12280
|
2236 |
statement becomes is included in the hypothesis, avoiding the
|
wenzelm@12280
|
2237 |
slightly cumbersome show "PROP ?case" form;
|
wenzelm@12280
|
2238 |
- 'induct' may now use elim-style induction rules without chaining
|
wenzelm@12280
|
2239 |
facts, using ``missing'' premises from the goal state; this allows
|
wenzelm@12280
|
2240 |
rules stemming from inductive sets to be applied in unstructured
|
wenzelm@12280
|
2241 |
scripts, while still benefitting from proper handling of non-atomic
|
wenzelm@12280
|
2242 |
statements; NB: major inductive premises need to be put first, all
|
wenzelm@12280
|
2243 |
the rest of the goal is passed through the induction;
|
wenzelm@12280
|
2244 |
- 'induct' proper support for mutual induction involving non-atomic
|
wenzelm@12280
|
2245 |
rule statements (uses the new concept of simultaneous goals, see
|
wenzelm@12280
|
2246 |
below);
|
wenzelm@12853
|
2247 |
- append all possible rule selections, but only use the first
|
wenzelm@12853
|
2248 |
success (no backtracking);
|
wenzelm@11995
|
2249 |
- removed obsolete "(simplified)" and "(stripped)" options of methods;
|
wenzelm@12754
|
2250 |
- undeclared rule case names default to numbers 1, 2, 3, ...;
|
wenzelm@12754
|
2251 |
- added 'print_induct_rules' (covered by help item in recent Proof
|
wenzelm@12754
|
2252 |
General versions);
|
wenzelm@11995
|
2253 |
- moved induct/cases attributes to Pure, methods to Provers;
|
wenzelm@11995
|
2254 |
- generic method setup instantiated for FOL and HOL;
|
wenzelm@11986
|
2255 |
|
wenzelm@12163
|
2256 |
* Pure: support multiple simultaneous goal statements, for example
|
wenzelm@12163
|
2257 |
"have a: A and b: B" (same for 'theorem' etc.); being a pure
|
wenzelm@12163
|
2258 |
meta-level mechanism, this acts as if several individual goals had
|
wenzelm@12163
|
2259 |
been stated separately; in particular common proof methods need to be
|
wenzelm@12163
|
2260 |
repeated in order to cover all claims; note that a single elimination
|
wenzelm@12163
|
2261 |
step is *not* sufficient to establish the two conjunctions, so this
|
wenzelm@12163
|
2262 |
fails:
|
wenzelm@12163
|
2263 |
|
wenzelm@12163
|
2264 |
assume "A & B" then have A and B .. (*".." fails*)
|
wenzelm@12163
|
2265 |
|
wenzelm@12163
|
2266 |
better use "obtain" in situations as above; alternative refer to
|
wenzelm@12163
|
2267 |
multi-step methods like 'auto', 'simp_all', 'blast+' etc.;
|
wenzelm@12163
|
2268 |
|
wenzelm@12078
|
2269 |
* Pure: proper integration with ``locales''; unlike the original
|
webertj@15154
|
2270 |
version by Florian Kammller, Isar locales package high-level proof
|
wenzelm@12078
|
2271 |
contexts rather than raw logical ones (e.g. we admit to include
|
wenzelm@12280
|
2272 |
attributes everywhere); operations on locales include merge and
|
wenzelm@12964
|
2273 |
rename; support for implicit arguments (``structures''); simultaneous
|
wenzelm@12964
|
2274 |
type-inference over imports and text; see also HOL/ex/Locales.thy for
|
wenzelm@12964
|
2275 |
some examples;
|
wenzelm@12078
|
2276 |
|
wenzelm@12707
|
2277 |
* Pure: the following commands have been ``localized'', supporting a
|
wenzelm@12707
|
2278 |
target locale specification "(in name)": 'lemma', 'theorem',
|
wenzelm@12707
|
2279 |
'corollary', 'lemmas', 'theorems', 'declare'; the results will be
|
wenzelm@12707
|
2280 |
stored both within the locale and at the theory level (exported and
|
wenzelm@12707
|
2281 |
qualified by the locale name);
|
wenzelm@12707
|
2282 |
|
wenzelm@12964
|
2283 |
* Pure: theory goals may now be specified in ``long'' form, with
|
wenzelm@12964
|
2284 |
ad-hoc contexts consisting of arbitrary locale elements. for example
|
wenzelm@12964
|
2285 |
``lemma foo: fixes x assumes "A x" shows "B x"'' (local syntax and
|
wenzelm@12964
|
2286 |
definitions may be given, too); the result is a meta-level rule with
|
wenzelm@12964
|
2287 |
the context elements being discharged in the obvious way;
|
wenzelm@12964
|
2288 |
|
wenzelm@12964
|
2289 |
* Pure: new proof command 'using' allows to augment currently used
|
wenzelm@12964
|
2290 |
facts after a goal statement ('using' is syntactically analogous to
|
wenzelm@12964
|
2291 |
'apply', but acts on the goal's facts only); this allows chained facts
|
wenzelm@12964
|
2292 |
to be separated into parts given before and after a claim, as in
|
wenzelm@12964
|
2293 |
``from a and b have C using d and e <proof>'';
|
wenzelm@12078
|
2294 |
|
wenzelm@11722
|
2295 |
* Pure: renamed "antecedent" case to "rule_context";
|
wenzelm@11722
|
2296 |
|
wenzelm@12964
|
2297 |
* Pure: new 'judgment' command records explicit information about the
|
wenzelm@12964
|
2298 |
object-logic embedding (used by several tools internally); no longer
|
wenzelm@12964
|
2299 |
use hard-wired "Trueprop";
|
wenzelm@12964
|
2300 |
|
wenzelm@11738
|
2301 |
* Pure: added 'corollary' command;
|
wenzelm@11738
|
2302 |
|
wenzelm@11722
|
2303 |
* Pure: fixed 'token_translation' command;
|
wenzelm@11722
|
2304 |
|
wenzelm@11899
|
2305 |
* Pure: removed obsolete 'exported' attribute;
|
wenzelm@11899
|
2306 |
|
wenzelm@11933
|
2307 |
* Pure: dummy pattern "_" in is/let is now automatically lifted over
|
wenzelm@11933
|
2308 |
bound variables: "ALL x. P x --> Q x" (is "ALL x. _ --> ?C x")
|
wenzelm@11899
|
2309 |
supersedes more cumbersome ... (is "ALL x. _ x --> ?C x");
|
wenzelm@11899
|
2310 |
|
wenzelm@11952
|
2311 |
* Pure: method 'atomize' presents local goal premises as object-level
|
wenzelm@11952
|
2312 |
statements (atomic meta-level propositions); setup controlled via
|
wenzelm@11952
|
2313 |
rewrite rules declarations of 'atomize' attribute; example
|
wenzelm@11952
|
2314 |
application: 'induct' method with proper rule statements in improper
|
wenzelm@11952
|
2315 |
proof *scripts*;
|
wenzelm@11952
|
2316 |
|
wenzelm@12106
|
2317 |
* Pure: emulation of instantiation tactics (rule_tac, cut_tac, etc.)
|
wenzelm@12106
|
2318 |
now consider the syntactic context of assumptions, giving a better
|
wenzelm@12106
|
2319 |
chance to get type-inference of the arguments right (this is
|
wenzelm@12106
|
2320 |
especially important for locales);
|
wenzelm@12106
|
2321 |
|
wenzelm@12312
|
2322 |
* Pure: "sorry" no longer requires quick_and_dirty in interactive
|
wenzelm@12312
|
2323 |
mode;
|
wenzelm@12312
|
2324 |
|
wenzelm@12405
|
2325 |
* Pure/obtain: the formal conclusion "thesis", being marked as
|
wenzelm@12405
|
2326 |
``internal'', may no longer be reference directly in the text;
|
wenzelm@12405
|
2327 |
potential INCOMPATIBILITY, may need to use "?thesis" in rare
|
wenzelm@12405
|
2328 |
situations;
|
wenzelm@12405
|
2329 |
|
wenzelm@12405
|
2330 |
* Pure: generic 'sym' attribute which declares a rule both as pure
|
wenzelm@12405
|
2331 |
'elim?' and for the 'symmetric' operation;
|
wenzelm@12405
|
2332 |
|
wenzelm@12877
|
2333 |
* Pure: marginal comments ``--'' may now occur just anywhere in the
|
wenzelm@12877
|
2334 |
text; the fixed correlation with particular command syntax has been
|
wenzelm@12877
|
2335 |
discontinued;
|
wenzelm@12877
|
2336 |
|
berghofe@13023
|
2337 |
* Pure: new method 'rules' is particularly well-suited for proof
|
berghofe@13023
|
2338 |
search in intuitionistic logic; a bit slower than 'blast' or 'fast',
|
berghofe@13023
|
2339 |
but often produces more compact proof terms with less detours;
|
berghofe@13023
|
2340 |
|
wenzelm@12364
|
2341 |
* Pure/Provers/classical: simplified integration with pure rule
|
wenzelm@12364
|
2342 |
attributes and methods; the classical "intro?/elim?/dest?"
|
wenzelm@12364
|
2343 |
declarations coincide with the pure ones; the "rule" method no longer
|
wenzelm@12364
|
2344 |
includes classically swapped intros; "intro" and "elim" methods no
|
wenzelm@12364
|
2345 |
longer pick rules from the context; also got rid of ML declarations
|
wenzelm@12364
|
2346 |
AddXIs/AddXEs/AddXDs; all of this has some potential for
|
wenzelm@12364
|
2347 |
INCOMPATIBILITY;
|
wenzelm@12364
|
2348 |
|
wenzelm@12405
|
2349 |
* Provers/classical: attribute 'swapped' produces classical inversions
|
wenzelm@12405
|
2350 |
of introduction rules;
|
wenzelm@12405
|
2351 |
|
wenzelm@12364
|
2352 |
* Provers/simplifier: 'simplified' attribute may refer to explicit
|
wenzelm@12364
|
2353 |
rules instead of full simplifier context; 'iff' attribute handles
|
wenzelm@12364
|
2354 |
conditional rules;
|
wenzelm@11936
|
2355 |
|
wenzelm@11745
|
2356 |
* HOL: 'typedef' now allows alternative names for Rep/Abs morphisms;
|
wenzelm@11745
|
2357 |
|
wenzelm@11690
|
2358 |
* HOL: 'recdef' now fails on unfinished automated proofs, use
|
wenzelm@11633
|
2359 |
"(permissive)" option to recover old behavior;
|
wenzelm@11633
|
2360 |
|
wenzelm@11933
|
2361 |
* HOL: 'inductive' no longer features separate (collective) attributes
|
wenzelm@11933
|
2362 |
for 'intros' (was found too confusing);
|
wenzelm@11933
|
2363 |
|
wenzelm@12405
|
2364 |
* HOL: properly declared induction rules less_induct and
|
wenzelm@12405
|
2365 |
wf_induct_rule;
|
wenzelm@12405
|
2366 |
|
kleing@11788
|
2367 |
|
wenzelm@11474
|
2368 |
*** HOL ***
|
wenzelm@11474
|
2369 |
|
wenzelm@11702
|
2370 |
* HOL: moved over to sane numeral syntax; the new policy is as
|
wenzelm@11702
|
2371 |
follows:
|
wenzelm@11702
|
2372 |
|
wenzelm@11702
|
2373 |
- 0 and 1 are polymorphic constants, which are defined on any
|
wenzelm@11702
|
2374 |
numeric type (nat, int, real etc.);
|
wenzelm@11702
|
2375 |
|
wenzelm@11702
|
2376 |
- 2, 3, 4, ... and -1, -2, -3, ... are polymorphic numerals, based
|
wenzelm@11702
|
2377 |
binary representation internally;
|
wenzelm@11702
|
2378 |
|
wenzelm@11702
|
2379 |
- type nat has special constructor Suc, and generally prefers Suc 0
|
wenzelm@11702
|
2380 |
over 1::nat and Suc (Suc 0) over 2::nat;
|
wenzelm@11702
|
2381 |
|
wenzelm@12364
|
2382 |
This change may cause significant problems of INCOMPATIBILITY; here
|
wenzelm@12364
|
2383 |
are some hints on converting existing sources:
|
wenzelm@11702
|
2384 |
|
wenzelm@11702
|
2385 |
- due to the new "num" token, "-0" and "-1" etc. are now atomic
|
wenzelm@11702
|
2386 |
entities, so expressions involving "-" (unary or binary minus) need
|
wenzelm@11702
|
2387 |
to be spaced properly;
|
wenzelm@11702
|
2388 |
|
wenzelm@11702
|
2389 |
- existing occurrences of "1" may need to be constraint "1::nat" or
|
wenzelm@11702
|
2390 |
even replaced by Suc 0; similar for old "2";
|
wenzelm@11702
|
2391 |
|
wenzelm@11702
|
2392 |
- replace "#nnn" by "nnn", and "#-nnn" by "-nnn";
|
wenzelm@11702
|
2393 |
|
wenzelm@11702
|
2394 |
- remove all special provisions on numerals in proofs;
|
wenzelm@11702
|
2395 |
|
wenzelm@13042
|
2396 |
* HOL: simp rules nat_number expand numerals on nat to Suc/0
|
wenzelm@12837
|
2397 |
representation (depends on bin_arith_simps in the default context);
|
wenzelm@12837
|
2398 |
|
wenzelm@12736
|
2399 |
* HOL: symbolic syntax for x^2 (numeral 2);
|
wenzelm@12736
|
2400 |
|
wenzelm@12335
|
2401 |
* HOL: the class of all HOL types is now called "type" rather than
|
wenzelm@12335
|
2402 |
"term"; INCOMPATIBILITY, need to adapt references to this type class
|
wenzelm@12335
|
2403 |
in axclass/classes, instance/arities, and (usually rare) occurrences
|
wenzelm@12335
|
2404 |
in typings (of consts etc.); internally the class is called
|
wenzelm@12335
|
2405 |
"HOL.type", ML programs should refer to HOLogic.typeS;
|
wenzelm@12335
|
2406 |
|
wenzelm@12280
|
2407 |
* HOL/record package improvements:
|
wenzelm@12280
|
2408 |
- new derived operations "fields" to build a partial record section,
|
wenzelm@12280
|
2409 |
"extend" to promote a fixed record to a record scheme, and
|
wenzelm@12280
|
2410 |
"truncate" for the reverse; cf. theorems "xxx.defs", which are *not*
|
wenzelm@12280
|
2411 |
declared as simp by default;
|
wenzelm@12587
|
2412 |
- shared operations ("more", "fields", etc.) now need to be always
|
wenzelm@12587
|
2413 |
qualified) --- potential INCOMPATIBILITY;
|
wenzelm@12280
|
2414 |
- removed "make_scheme" operations (use "make" with "extend") --
|
wenzelm@12280
|
2415 |
INCOMPATIBILITY;
|
wenzelm@11937
|
2416 |
- removed "more" class (simply use "term") -- INCOMPATIBILITY;
|
wenzelm@12253
|
2417 |
- provides cases/induct rules for use with corresponding Isar
|
wenzelm@12253
|
2418 |
methods (for concrete records, record schemes, concrete more
|
wenzelm@12280
|
2419 |
parts, and schematic more parts -- in that order);
|
wenzelm@11930
|
2420 |
- internal definitions directly based on a light-weight abstract
|
wenzelm@11930
|
2421 |
theory of product types over typedef rather than datatype;
|
wenzelm@11930
|
2422 |
|
berghofe@13023
|
2423 |
* HOL: generic code generator for generating executable ML code from
|
berghofe@13023
|
2424 |
specifications; specific support for HOL constructs such as inductive
|
berghofe@13023
|
2425 |
datatypes and sets, as well as recursive functions; can be invoked
|
berghofe@13023
|
2426 |
via 'generate_code' theory section;
|
berghofe@13023
|
2427 |
|
wenzelm@11933
|
2428 |
* HOL: canonical cases/induct rules for n-tuples (n = 3..7);
|
wenzelm@11933
|
2429 |
|
paulson@13824
|
2430 |
* HOL: consolidated and renamed several theories. In particular:
|
wenzelm@14731
|
2431 |
Ord.thy has been absorbed into HOL.thy
|
wenzelm@14731
|
2432 |
String.thy has been absorbed into List.thy
|
wenzelm@14731
|
2433 |
|
wenzelm@11802
|
2434 |
* HOL: concrete setsum syntax "\<Sum>i:A. b" == "setsum (%i. b) A"
|
wenzelm@11802
|
2435 |
(beware of argument permutation!);
|
wenzelm@11802
|
2436 |
|
wenzelm@11657
|
2437 |
* HOL: linorder_less_split superseded by linorder_cases;
|
wenzelm@11657
|
2438 |
|
wenzelm@12917
|
2439 |
* HOL/List: "nodups" renamed to "distinct";
|
nipkow@12889
|
2440 |
|
wenzelm@11633
|
2441 |
* HOL: added "The" definite description operator; move Hilbert's "Eps"
|
paulson@13824
|
2442 |
to peripheral theory "Hilbert_Choice"; some INCOMPATIBILITIES:
|
paulson@13824
|
2443 |
- Ex_def has changed, now need to use some_eq_ex
|
wenzelm@11437
|
2444 |
|
wenzelm@11572
|
2445 |
* HOL: made split_all_tac safe; EXISTING PROOFS MAY FAIL OR LOOP, so
|
wenzelm@11572
|
2446 |
in this (rare) case use:
|
wenzelm@11572
|
2447 |
|
wenzelm@11572
|
2448 |
delSWrapper "split_all_tac"
|
wenzelm@11572
|
2449 |
addSbefore ("unsafe_split_all_tac", unsafe_split_all_tac)
|
wenzelm@11572
|
2450 |
|
wenzelm@11572
|
2451 |
* HOL: added safe wrapper "split_conv_tac" to claset; EXISTING PROOFS
|
wenzelm@11474
|
2452 |
MAY FAIL;
|
nipkow@11361
|
2453 |
|
wenzelm@11572
|
2454 |
* HOL: introduced f^n = f o ... o f; warning: due to the limits of
|
wenzelm@11572
|
2455 |
Isabelle's type classes, ^ on functions and relations has too general
|
wenzelm@11572
|
2456 |
a domain, namely ('a * 'b) set and 'a => 'b; this means that it may be
|
wenzelm@11572
|
2457 |
necessary to attach explicit type constraints;
|
nipkow@11307
|
2458 |
|
wenzelm@12917
|
2459 |
* HOL/Relation: the prefix name of the infix "O" has been changed from
|
wenzelm@12917
|
2460 |
"comp" to "rel_comp"; INCOMPATIBILITY: a few theorems have been
|
wenzelm@12917
|
2461 |
renamed accordingly (eg "compI" -> "rel_compI").
|
nipkow@12489
|
2462 |
|
wenzelm@11487
|
2463 |
* HOL: syntax translations now work properly with numerals and records
|
wenzelm@11487
|
2464 |
expressions;
|
wenzelm@11474
|
2465 |
|
wenzelm@12457
|
2466 |
* HOL: bounded abstraction now uses syntax "%" / "\<lambda>" instead
|
wenzelm@12457
|
2467 |
of "lam" -- INCOMPATIBILITY;
|
wenzelm@11474
|
2468 |
|
wenzelm@11933
|
2469 |
* HOL: got rid of some global declarations (potential INCOMPATIBILITY
|
wenzelm@11933
|
2470 |
for ML tools): const "()" renamed "Product_Type.Unity", type "unit"
|
wenzelm@11933
|
2471 |
renamed "Product_Type.unit";
|
wenzelm@11611
|
2472 |
|
nipkow@12564
|
2473 |
* HOL: renamed rtrancl_into_rtrancl2 to converse_rtrancl_into_rtrancl
|
nipkow@12564
|
2474 |
|
wenzelm@12924
|
2475 |
* HOL: removed obsolete theorem "optionE" (use "option.exhaust", or
|
wenzelm@12924
|
2476 |
the "cases" method);
|
wenzelm@12924
|
2477 |
|
wenzelm@12597
|
2478 |
* HOL/GroupTheory: group theory examples including Sylow's theorem (by
|
webertj@15154
|
2479 |
Florian Kammller);
|
wenzelm@12597
|
2480 |
|
wenzelm@12608
|
2481 |
* HOL/IMP: updated and converted to new-style theory format; several
|
wenzelm@12608
|
2482 |
parts turned into readable document, with proper Isar proof texts and
|
wenzelm@12608
|
2483 |
some explanations (by Gerwin Klein);
|
wenzelm@12597
|
2484 |
|
wenzelm@12734
|
2485 |
* HOL-Real: added Complex_Numbers (by Gertrud Bauer);
|
wenzelm@12734
|
2486 |
|
wenzelm@12690
|
2487 |
* HOL-Hyperreal is now a logic image;
|
wenzelm@12690
|
2488 |
|
wenzelm@11611
|
2489 |
|
wenzelm@12022
|
2490 |
*** HOLCF ***
|
wenzelm@12022
|
2491 |
|
wenzelm@12622
|
2492 |
* Isar: consts/constdefs supports mixfix syntax for continuous
|
wenzelm@12622
|
2493 |
operations;
|
wenzelm@12622
|
2494 |
|
wenzelm@12622
|
2495 |
* Isar: domain package adapted to new-style theory format, e.g. see
|
wenzelm@12622
|
2496 |
HOLCF/ex/Dnat.thy;
|
wenzelm@12622
|
2497 |
|
wenzelm@12622
|
2498 |
* theory Lift: proper use of rep_datatype lift instead of ML hacks --
|
wenzelm@12280
|
2499 |
potential INCOMPATIBILITY; now use plain induct_tac instead of former
|
wenzelm@12280
|
2500 |
lift.induct_tac, always use UU instead of Undef;
|
wenzelm@12022
|
2501 |
|
wenzelm@12597
|
2502 |
* HOLCF/IMP: updated and converted to new-style theory;
|
wenzelm@12597
|
2503 |
|
wenzelm@12022
|
2504 |
|
wenzelm@11474
|
2505 |
*** ZF ***
|
wenzelm@11474
|
2506 |
|
wenzelm@12622
|
2507 |
* Isar: proper integration of logic-specific tools and packages,
|
wenzelm@12622
|
2508 |
including theory commands '(co)inductive', '(co)datatype',
|
wenzelm@12622
|
2509 |
'rep_datatype', 'inductive_cases', as well as methods 'ind_cases',
|
wenzelm@12622
|
2510 |
'induct_tac', 'case_tac', and 'typecheck' (with attribute 'TC');
|
wenzelm@12622
|
2511 |
|
wenzelm@12622
|
2512 |
* theory Main no longer includes AC; for the Axiom of Choice, base
|
wenzelm@12622
|
2513 |
your theory on Main_ZFC;
|
wenzelm@12622
|
2514 |
|
wenzelm@12622
|
2515 |
* the integer library now covers quotients and remainders, with many
|
wenzelm@12622
|
2516 |
laws relating division to addition, multiplication, etc.;
|
paulson@12563
|
2517 |
|
wenzelm@12280
|
2518 |
* ZF/UNITY: Chandy and Misra's UNITY is now available in ZF, giving a
|
wenzelm@12280
|
2519 |
typeless version of the formalism;
|
wenzelm@12280
|
2520 |
|
wenzelm@13025
|
2521 |
* ZF/AC, Coind, IMP, Resid: updated and converted to new-style theory
|
wenzelm@13025
|
2522 |
format;
|
wenzelm@12608
|
2523 |
|
wenzelm@12280
|
2524 |
* ZF/Induct: new directory for examples of inductive definitions,
|
wenzelm@12608
|
2525 |
including theory Multiset for multiset orderings; converted to
|
wenzelm@12608
|
2526 |
new-style theory format;
|
wenzelm@12177
|
2527 |
|
wenzelm@13025
|
2528 |
* ZF: many new theorems about lists, ordinals, etc.;
|
paulson@12850
|
2529 |
|
wenzelm@11474
|
2530 |
|
wenzelm@11474
|
2531 |
*** General ***
|
wenzelm@11474
|
2532 |
|
wenzelm@12280
|
2533 |
* Pure/kernel: meta-level proof terms (by Stefan Berghofer); reference
|
wenzelm@12280
|
2534 |
variable proof controls level of detail: 0 = no proofs (only oracle
|
wenzelm@12280
|
2535 |
dependencies), 1 = lemma dependencies, 2 = compact proof terms; see
|
wenzelm@12280
|
2536 |
also ref manual for further ML interfaces;
|
wenzelm@12280
|
2537 |
|
wenzelm@12280
|
2538 |
* Pure/axclass: removed obsolete ML interface
|
wenzelm@12280
|
2539 |
goal_subclass/goal_arity;
|
wenzelm@12280
|
2540 |
|
wenzelm@12280
|
2541 |
* Pure/syntax: new token syntax "num" for plain numerals (without "#"
|
wenzelm@12280
|
2542 |
of "xnum"); potential INCOMPATIBILITY, since -0, -1 etc. are now
|
wenzelm@12280
|
2543 |
separate tokens, so expressions involving minus need to be spaced
|
wenzelm@12280
|
2544 |
properly;
|
wenzelm@12280
|
2545 |
|
wenzelm@12312
|
2546 |