wenzelm@57491
|
1 |
Isabelle NEWS -- history of user-relevant changes
|
wenzelm@57491
|
2 |
=================================================
|
wenzelm@2553
|
3 |
|
wenzelm@60006
|
4 |
(Note: Isabelle/jEdit shows a tree-view of this file in Sidekick.)
|
wenzelm@60006
|
5 |
|
wenzelm@60331
|
6 |
|
wenzelm@60138
|
7 |
New in this Isabelle version
|
wenzelm@60138
|
8 |
----------------------------
|
wenzelm@60138
|
9 |
|
wenzelm@61337
|
10 |
*** General ***
|
wenzelm@61337
|
11 |
|
wenzelm@61337
|
12 |
* Toplevel theorem statements have been simplified as follows:
|
wenzelm@61337
|
13 |
|
wenzelm@61337
|
14 |
theorems ~> lemmas
|
wenzelm@61337
|
15 |
schematic_lemma ~> schematic_goal
|
wenzelm@61337
|
16 |
schematic_theorem ~> schematic_goal
|
wenzelm@61337
|
17 |
schematic_corollary ~> schematic_goal
|
wenzelm@61337
|
18 |
|
wenzelm@61337
|
19 |
Command-line tool "isabelle update_theorems" updates theory sources
|
wenzelm@61337
|
20 |
accordingly.
|
wenzelm@61337
|
21 |
|
wenzelm@61338
|
22 |
* Toplevel theorem statement 'proposition' is another alias for
|
wenzelm@61338
|
23 |
'theorem'.
|
wenzelm@61338
|
24 |
|
wenzelm@61380
|
25 |
* HTML presentation uses the standard IsabelleText font and Unicode
|
wenzelm@61380
|
26 |
rendering of Isabelle symbols like Isabelle/Scala/jEdit. The former
|
wenzelm@61380
|
27 |
print mode "HTML" looses its special meaning.
|
wenzelm@61380
|
28 |
|
wenzelm@61337
|
29 |
|
wenzelm@60610
|
30 |
*** Prover IDE -- Isabelle/Scala/jEdit ***
|
wenzelm@60610
|
31 |
|
wenzelm@60610
|
32 |
* Improved scheduling for urgent print tasks (e.g. command state output,
|
wenzelm@60610
|
33 |
interactive queries) wrt. long-running background tasks.
|
wenzelm@60610
|
34 |
|
wenzelm@60986
|
35 |
* IDE support for the source-level debugger of Poly/ML, to work with
|
wenzelm@60984
|
36 |
Isabelle/ML and official Standard ML. Configuration option "ML_debugger"
|
wenzelm@60984
|
37 |
and commands 'ML_file_debug', 'ML_file_no_debug', 'SML_file_debug',
|
wenzelm@60984
|
38 |
'SML_file_no_debug' control compilation of sources with debugging
|
wenzelm@60984
|
39 |
information. The Debugger panel allows to set breakpoints (via context
|
wenzelm@60984
|
40 |
menu), step through stopped threads, evaluate local ML expressions etc.
|
wenzelm@60984
|
41 |
At least one Debugger view needs to be active to have any effect on the
|
wenzelm@60984
|
42 |
running ML program.
|
wenzelm@60984
|
43 |
|
wenzelm@61173
|
44 |
* The main Isabelle executable is managed as single-instance Desktop
|
wenzelm@61173
|
45 |
application uniformly on all platforms: Linux, Windows, Mac OS X.
|
wenzelm@61170
|
46 |
|
wenzelm@61198
|
47 |
* The text overview column (status of errors, warnings etc.) is updated
|
wenzelm@61198
|
48 |
asynchronously, leading to much better editor reactivity. Moreover, the
|
wenzelm@61198
|
49 |
full document node content is taken into account.
|
wenzelm@61198
|
50 |
|
wenzelm@61218
|
51 |
* The State panel manages explicit proof state output, with jEdit action
|
wenzelm@61218
|
52 |
"isabelle.update-state" (shortcut S+ENTER) to trigger update according
|
wenzelm@61218
|
53 |
to cursor position. Option "editor_output_state" controls implicit proof
|
wenzelm@61218
|
54 |
state output in the Output panel: suppressing this reduces resource
|
wenzelm@61218
|
55 |
requirements of prover time and GUI space.
|
wenzelm@61215
|
56 |
|
wenzelm@60610
|
57 |
|
wenzelm@61405
|
58 |
*** Document preparation ***
|
wenzelm@61405
|
59 |
|
wenzelm@61405
|
60 |
* Isabelle control symbols for markup and formatting:
|
wenzelm@61405
|
61 |
|
wenzelm@61405
|
62 |
\<^smallskip> \smallskip
|
wenzelm@61405
|
63 |
\<^medskip> \medskip
|
wenzelm@61405
|
64 |
\<^bigskip> \bigskip
|
wenzelm@61405
|
65 |
|
wenzelm@61405
|
66 |
\<^item> \item (itemize)
|
wenzelm@61405
|
67 |
\<^enum> \item (enumeration)
|
wenzelm@61405
|
68 |
|
wenzelm@61405
|
69 |
|
wenzelm@60406
|
70 |
*** Isar ***
|
wenzelm@60406
|
71 |
|
wenzelm@60406
|
72 |
* Command 'obtain' binds term abbreviations (via 'is' patterns) in the
|
wenzelm@60406
|
73 |
proof body as well, abstracted over relevant parameters.
|
wenzelm@60406
|
74 |
|
wenzelm@60477
|
75 |
* Improved type-inference for theorem statement 'obtains': separate
|
wenzelm@60477
|
76 |
parameter scope for of each clause.
|
wenzelm@60477
|
77 |
|
wenzelm@60408
|
78 |
* Term abbreviations via 'is' patterns also work for schematic
|
wenzelm@60408
|
79 |
statements: result is abstracted over unknowns.
|
wenzelm@60408
|
80 |
|
wenzelm@60414
|
81 |
* Local goals ('have', 'show', 'hence', 'thus') allow structured
|
wenzelm@60414
|
82 |
statements like fixes/assumes/shows in theorem specifications, but the
|
wenzelm@60555
|
83 |
notation is postfix with keywords 'if' (or 'when') and 'for'. For
|
wenzelm@60555
|
84 |
example:
|
wenzelm@60414
|
85 |
|
wenzelm@60414
|
86 |
have result: "C x y"
|
wenzelm@60414
|
87 |
if "A x" and "B y"
|
wenzelm@60414
|
88 |
for x :: 'a and y :: 'a
|
wenzelm@60414
|
89 |
<proof>
|
wenzelm@60414
|
90 |
|
wenzelm@60449
|
91 |
The local assumptions are bound to the name "that". The result is
|
wenzelm@60449
|
92 |
exported from context of the statement as usual. The above roughly
|
wenzelm@60414
|
93 |
corresponds to a raw proof block like this:
|
wenzelm@60414
|
94 |
|
wenzelm@60414
|
95 |
{
|
wenzelm@60414
|
96 |
fix x :: 'a and y :: 'a
|
wenzelm@60449
|
97 |
assume that: "A x" "B y"
|
wenzelm@60414
|
98 |
have "C x y" <proof>
|
wenzelm@60414
|
99 |
}
|
wenzelm@60414
|
100 |
note result = this
|
wenzelm@60406
|
101 |
|
wenzelm@60555
|
102 |
The keyword 'when' may be used instead of 'if', to indicate 'presume'
|
wenzelm@60555
|
103 |
instead of 'assume' above.
|
wenzelm@60555
|
104 |
|
wenzelm@60595
|
105 |
* The meaning of 'show' with Pure rule statements has changed: premises
|
wenzelm@60595
|
106 |
are treated in the sense of 'assume', instead of 'presume'. This means,
|
wenzelm@60595
|
107 |
a goal like "\<And>x. A x \<Longrightarrow> B x \<Longrightarrow> C x" can be solved completely as follows:
|
wenzelm@60595
|
108 |
|
wenzelm@60595
|
109 |
show "\<And>x. A x \<Longrightarrow> B x \<Longrightarrow> C x"
|
wenzelm@60595
|
110 |
|
wenzelm@60595
|
111 |
or:
|
wenzelm@60595
|
112 |
|
wenzelm@60595
|
113 |
show "C x" if "A x" "B x" for x
|
wenzelm@60595
|
114 |
|
wenzelm@60595
|
115 |
Rare INCOMPATIBILITY, the old behaviour may be recovered as follows:
|
wenzelm@60595
|
116 |
|
wenzelm@60595
|
117 |
show "C x" when "A x" "B x" for x
|
wenzelm@60595
|
118 |
|
wenzelm@60406
|
119 |
* New command 'supply' supports fact definitions during goal refinement
|
wenzelm@60406
|
120 |
('apply' scripts).
|
wenzelm@60406
|
121 |
|
wenzelm@60459
|
122 |
* New command 'consider' states rules for generalized elimination and
|
wenzelm@60459
|
123 |
case splitting. This is like a toplevel statement "theorem obtains" used
|
wenzelm@60459
|
124 |
within a proof body; or like a multi-branch 'obtain' without activation
|
wenzelm@60459
|
125 |
of the local context elements yet.
|
wenzelm@60459
|
126 |
|
wenzelm@60455
|
127 |
* Proof method "cases" allows to specify the rule as first entry of
|
wenzelm@60455
|
128 |
chained facts. This is particularly useful with 'consider':
|
wenzelm@60455
|
129 |
|
wenzelm@60455
|
130 |
consider (a) A | (b) B | (c) C <proof>
|
wenzelm@60455
|
131 |
then have something
|
wenzelm@60455
|
132 |
proof cases
|
wenzelm@60455
|
133 |
case a
|
wenzelm@60455
|
134 |
then show ?thesis <proof>
|
wenzelm@60455
|
135 |
next
|
wenzelm@60455
|
136 |
case b
|
wenzelm@60455
|
137 |
then show ?thesis <proof>
|
wenzelm@60455
|
138 |
next
|
wenzelm@60455
|
139 |
case c
|
wenzelm@60455
|
140 |
then show ?thesis <proof>
|
wenzelm@60455
|
141 |
qed
|
wenzelm@60455
|
142 |
|
wenzelm@60565
|
143 |
* Command 'case' allows fact name and attribute specification like this:
|
wenzelm@60565
|
144 |
|
wenzelm@60565
|
145 |
case a: (c xs)
|
wenzelm@60565
|
146 |
case a [attributes]: (c xs)
|
wenzelm@60565
|
147 |
|
wenzelm@60565
|
148 |
Facts that are introduced by invoking the case context are uniformly
|
wenzelm@60565
|
149 |
qualified by "a"; the same name is used for the cumulative fact. The old
|
wenzelm@60565
|
150 |
form "case (c xs) [attributes]" is no longer supported. Rare
|
wenzelm@60565
|
151 |
INCOMPATIBILITY, need to adapt uses of case facts in exotic situations,
|
wenzelm@60565
|
152 |
and always put attributes in front.
|
wenzelm@60565
|
153 |
|
wenzelm@60618
|
154 |
* The standard proof method of commands 'proof' and '..' is now called
|
wenzelm@60618
|
155 |
"standard" to make semantically clear what it is; the old name "default"
|
wenzelm@60618
|
156 |
is still available as legacy for some time. Documentation now explains
|
wenzelm@60618
|
157 |
'..' more accurately as "by standard" instead of "by rule".
|
wenzelm@60618
|
158 |
|
wenzelm@60631
|
159 |
* Command 'subgoal' allows to impose some structure on backward
|
wenzelm@60631
|
160 |
refinements, to avoid proof scripts degenerating into long of 'apply'
|
wenzelm@60631
|
161 |
sequences. Further explanations and examples are given in the isar-ref
|
wenzelm@60631
|
162 |
manual.
|
wenzelm@60631
|
163 |
|
wenzelm@61166
|
164 |
* Proof method "goal_cases" turns the current subgoals into cases within
|
wenzelm@61166
|
165 |
the context; the conclusion is bound to variable ?case in each case. For
|
wenzelm@61166
|
166 |
example:
|
wenzelm@60617
|
167 |
|
wenzelm@60617
|
168 |
lemma "\<And>x. A x \<Longrightarrow> B x \<Longrightarrow> C x"
|
wenzelm@60622
|
169 |
and "\<And>y z. U y \<Longrightarrow> V z \<Longrightarrow> W y z"
|
wenzelm@61166
|
170 |
proof goal_cases
|
wenzelm@60622
|
171 |
case (1 x)
|
wenzelm@60622
|
172 |
then show ?case using \<open>A x\<close> \<open>B x\<close> sorry
|
wenzelm@60622
|
173 |
next
|
wenzelm@60622
|
174 |
case (2 y z)
|
wenzelm@60622
|
175 |
then show ?case using \<open>U y\<close> \<open>V z\<close> sorry
|
wenzelm@60622
|
176 |
qed
|
wenzelm@60622
|
177 |
|
wenzelm@60622
|
178 |
lemma "\<And>x. A x \<Longrightarrow> B x \<Longrightarrow> C x"
|
wenzelm@60622
|
179 |
and "\<And>y z. U y \<Longrightarrow> V z \<Longrightarrow> W y z"
|
wenzelm@61166
|
180 |
proof goal_cases
|
wenzelm@60617
|
181 |
case prems: 1
|
wenzelm@60617
|
182 |
then show ?case using prems sorry
|
wenzelm@60617
|
183 |
next
|
wenzelm@60617
|
184 |
case prems: 2
|
wenzelm@60617
|
185 |
then show ?case using prems sorry
|
wenzelm@60617
|
186 |
qed
|
wenzelm@60578
|
187 |
|
wenzelm@60581
|
188 |
* The undocumented feature of implicit cases goal1, goal2, goal3, etc.
|
wenzelm@60617
|
189 |
is marked as legacy, and will be removed eventually. The proof method
|
wenzelm@60617
|
190 |
"goals" achieves a similar effect within regular Isar; often it can be
|
wenzelm@60617
|
191 |
done more adequately by other means (e.g. 'consider').
|
wenzelm@60581
|
192 |
|
wenzelm@60551
|
193 |
* Nesting of Isar goal structure has been clarified: the context after
|
wenzelm@60551
|
194 |
the initial backwards refinement is retained for the whole proof, within
|
wenzelm@60551
|
195 |
all its context sections (as indicated via 'next'). This is e.g.
|
wenzelm@60551
|
196 |
relevant for 'using', 'including', 'supply':
|
wenzelm@60551
|
197 |
|
wenzelm@60551
|
198 |
have "A \<and> A" if a: A for A
|
wenzelm@60551
|
199 |
supply [simp] = a
|
wenzelm@60551
|
200 |
proof
|
wenzelm@60551
|
201 |
show A by simp
|
wenzelm@60551
|
202 |
next
|
wenzelm@60551
|
203 |
show A by simp
|
wenzelm@60551
|
204 |
qed
|
wenzelm@60551
|
205 |
|
wenzelm@60554
|
206 |
* Method "sleep" succeeds after a real-time delay (in seconds). This is
|
wenzelm@60554
|
207 |
occasionally useful for demonstration and testing purposes.
|
wenzelm@60554
|
208 |
|
wenzelm@60406
|
209 |
|
wenzelm@60331
|
210 |
*** Pure ***
|
wenzelm@60331
|
211 |
|
wenzelm@61252
|
212 |
* Command 'print_definitions' prints dependencies of definitional
|
wenzelm@61252
|
213 |
specifications. This functionality used to be part of 'print_theory'.
|
wenzelm@61252
|
214 |
|
wenzelm@60489
|
215 |
* The vacuous fact "TERM x" may be established "by fact" or as `TERM x`
|
wenzelm@60489
|
216 |
as well, not just "by this" or "." as before.
|
wenzelm@60489
|
217 |
|
wenzelm@60331
|
218 |
* Configuration option rule_insts_schematic has been discontinued
|
wenzelm@60331
|
219 |
(intermediate legacy feature in Isabelle2015). INCOMPATIBILITY.
|
wenzelm@60331
|
220 |
|
haftmann@60347
|
221 |
* Abbreviations in type classes now carry proper sort constraint.
|
haftmann@60347
|
222 |
Rare INCOMPATIBILITY in situations where the previous misbehaviour
|
haftmann@61228
|
223 |
has been exploited.
|
haftmann@60347
|
224 |
|
haftmann@60347
|
225 |
* Refinement of user-space type system in type classes: pseudo-local
|
haftmann@60347
|
226 |
operations behave more similar to abbreviations. Potential
|
haftmann@60347
|
227 |
INCOMPATIBILITY in exotic situations.
|
haftmann@60347
|
228 |
|
haftmann@60347
|
229 |
|
nipkow@60171
|
230 |
*** HOL ***
|
nipkow@60171
|
231 |
|
haftmann@61424
|
232 |
* Combinator to represent case distinction on products is named "case_prod",
|
haftmann@61424
|
233 |
uniformly, discontinuing any input aliasses. Very popular theorem aliasses
|
haftmann@61424
|
234 |
have been retained.
|
haftmann@61424
|
235 |
Consolidated facts:
|
haftmann@61424
|
236 |
PairE ~> prod.exhaust
|
haftmann@61424
|
237 |
Pair_eq ~> prod.inject
|
haftmann@61424
|
238 |
pair_collapse ~> prod.collapse
|
haftmann@61424
|
239 |
Pair_fst_snd_eq ~> prod_eq_iff
|
haftmann@61424
|
240 |
split_twice ~> prod.case_distrib
|
haftmann@61424
|
241 |
split_weak_cong ~> prod.case_cong_weak
|
haftmann@61424
|
242 |
split_split ~> prod.split
|
haftmann@61424
|
243 |
split_split_asm ~> prod.split_asm
|
haftmann@61424
|
244 |
splitI ~> case_prodI
|
haftmann@61424
|
245 |
splitD ~> case_prodD
|
haftmann@61424
|
246 |
splitI2 ~> case_prodI2
|
haftmann@61424
|
247 |
splitI2' ~> case_prodI2'
|
haftmann@61424
|
248 |
splitE ~> case_prodE
|
haftmann@61424
|
249 |
splitE' ~> case_prodE'
|
haftmann@61424
|
250 |
split_pair ~> case_prod_Pair
|
haftmann@61424
|
251 |
split_eta ~> case_prod_eta
|
haftmann@61424
|
252 |
split_comp ~> case_prod_comp
|
haftmann@61424
|
253 |
mem_splitI ~> mem_case_prodI
|
haftmann@61424
|
254 |
mem_splitI2 ~> mem_case_prodI2
|
haftmann@61424
|
255 |
mem_splitE ~> mem_case_prodE
|
haftmann@61424
|
256 |
The_split ~> The_case_prod
|
haftmann@61424
|
257 |
cond_split_eta ~> cond_case_prod_eta
|
haftmann@61424
|
258 |
Collect_split_in_rel_leE ~> Collect_case_prod_in_rel_leE
|
haftmann@61424
|
259 |
Collect_split_in_rel_leI ~> Collect_case_prod_in_rel_leI
|
haftmann@61424
|
260 |
in_rel_Collect_split_eq ~> in_rel_Collect_case_prod_eq
|
haftmann@61424
|
261 |
Collect_split_Grp_eqD ~> Collect_case_prod_Grp_eqD
|
haftmann@61424
|
262 |
Collect_split_Grp_inD ~> Collect_case_prod_Grp_in
|
haftmann@61424
|
263 |
Domain_Collect_split ~> Domain_Collect_case_prod
|
haftmann@61424
|
264 |
Image_Collect_split ~> Image_Collect_case_prod
|
haftmann@61424
|
265 |
Range_Collect_split ~> Range_Collect_case_prod
|
haftmann@61424
|
266 |
Eps_split ~> Eps_case_prod
|
haftmann@61424
|
267 |
Eps_split_eq ~> Eps_case_prod_eq
|
haftmann@61424
|
268 |
split_rsp ~> case_prod_rsp
|
haftmann@61424
|
269 |
curry_split ~> curry_case_prod
|
haftmann@61424
|
270 |
split_curry ~> case_prod_curry
|
haftmann@61424
|
271 |
Changes in structure HOLogic:
|
haftmann@61424
|
272 |
split_const ~> case_prod_const
|
haftmann@61424
|
273 |
mk_split ~> mk_case_prod
|
haftmann@61424
|
274 |
mk_psplits ~> mk_ptupleabs
|
haftmann@61424
|
275 |
strip_psplits ~> strip_ptupleabs
|
haftmann@61424
|
276 |
INCOMPATIBILITY.
|
haftmann@61424
|
277 |
|
wenzelm@61308
|
278 |
* Commands 'inductive' and 'inductive_set' work better when names for
|
wenzelm@61308
|
279 |
intro rules are omitted: the "cases" and "induct" rules no longer
|
wenzelm@61308
|
280 |
declare empty case_names, but no case_names at all. This allows to use
|
wenzelm@61308
|
281 |
numbered cases in proofs, without requiring method "goal_cases".
|
wenzelm@61308
|
282 |
|
wenzelm@61269
|
283 |
* The 'typedef' command has been upgraded from a partially checked
|
wenzelm@61269
|
284 |
"axiomatization", to a full definitional specification that takes the
|
wenzelm@61269
|
285 |
global collection of overloaded constant / type definitions into
|
wenzelm@61269
|
286 |
account. Type definitions with open dependencies on overloaded
|
wenzelm@61269
|
287 |
definitions need to be specified as "typedef (overloaded)". This
|
wenzelm@61269
|
288 |
provides extra robustness in theory construction. Rare INCOMPATIBILITY.
|
wenzelm@61269
|
289 |
|
wenzelm@61118
|
290 |
* Qualification of various formal entities in the libraries is done more
|
wenzelm@61118
|
291 |
uniformly via "context begin qualified definition ... end" instead of
|
wenzelm@61118
|
292 |
old-style "hide_const (open) ...". Consequently, both the defined
|
wenzelm@61118
|
293 |
constant and its defining fact become qualified, e.g. Option.is_none and
|
wenzelm@61118
|
294 |
Option.is_none_def. Occasional INCOMPATIBILITY in applications.
|
wenzelm@61118
|
295 |
|
wenzelm@61069
|
296 |
* Some old and rarely used ASCII replacement syntax has been removed.
|
wenzelm@61069
|
297 |
INCOMPATIBILITY, standard syntax with symbols should be used instead.
|
wenzelm@61069
|
298 |
The subsequent commands help to reproduce the old forms, e.g. to
|
wenzelm@61069
|
299 |
simplify porting old theories:
|
wenzelm@61069
|
300 |
|
wenzelm@61069
|
301 |
type_notation Map.map (infixr "~=>" 0)
|
wenzelm@61069
|
302 |
notation Map.map_comp (infixl "o'_m" 55)
|
wenzelm@61069
|
303 |
|
wenzelm@61384
|
304 |
type_notation FinFun.finfun ("(_ =>f /_)" [22, 21] 21)
|
wenzelm@61384
|
305 |
|
wenzelm@61384
|
306 |
notation FuncSet.funcset (infixr "->" 60)
|
wenzelm@61384
|
307 |
notation FuncSet.extensional_funcset (infixr "->\<^sub>E" 60)
|
wenzelm@61384
|
308 |
|
wenzelm@61384
|
309 |
notation Omega_Words_Fun.conc (infixr "conc" 65)
|
wenzelm@61384
|
310 |
|
wenzelm@61384
|
311 |
notation Preorder.equiv ("op ~~")
|
wenzelm@61384
|
312 |
and Preorder.equiv ("(_/ ~~ _)" [51, 51] 50)
|
wenzelm@61384
|
313 |
|
wenzelm@61143
|
314 |
* The alternative notation "\<Colon>" for type and sort constraints has been
|
wenzelm@61143
|
315 |
removed: in LaTeX document output it looks the same as "::".
|
wenzelm@61143
|
316 |
INCOMPATIBILITY, use plain "::" instead.
|
wenzelm@61143
|
317 |
|
wenzelm@60841
|
318 |
* Theory Map: lemma map_of_is_SomeD was a clone of map_of_SomeD and has
|
wenzelm@60841
|
319 |
been removed. INCOMPATIBILITY.
|
wenzelm@60841
|
320 |
|
lars@60712
|
321 |
* Quickcheck setup for finite sets.
|
lars@60712
|
322 |
|
nipkow@60171
|
323 |
* Discontinued simp_legacy_precond. Potential INCOMPATIBILITY.
|
wenzelm@60138
|
324 |
|
blanchet@60306
|
325 |
* Sledgehammer:
|
blanchet@61318
|
326 |
- The MaSh relevance filter has been sped up.
|
blanchet@60306
|
327 |
- Proof reconstruction has been improved, to minimize the incidence of
|
blanchet@60306
|
328 |
cases where Sledgehammer gives a proof that does not work.
|
blanchet@60306
|
329 |
- Auto Sledgehammer now minimizes and preplays the results.
|
blanchet@61030
|
330 |
- Handle Vampire 4.0 proof output without raising exception.
|
blanchet@61043
|
331 |
- Eliminated "MASH" environment variable. Use the "MaSh" option in
|
blanchet@61043
|
332 |
Isabelle/jEdit instead. INCOMPATIBILITY.
|
blanchet@61317
|
333 |
- Eliminated obsolete "blocking" option and related subcommands.
|
blanchet@60306
|
334 |
|
blanchet@60310
|
335 |
* Nitpick:
|
blanchet@61325
|
336 |
- Fixed soundness bug in translation of "finite" predicate.
|
blanchet@61324
|
337 |
- Fixed soundness bug in "destroy_constrs" optimization.
|
blanchet@60310
|
338 |
- Removed "check_potential" and "check_genuine" options.
|
blanchet@61317
|
339 |
- Eliminated obsolete "blocking" option.
|
blanchet@60310
|
340 |
|
blanchet@61345
|
341 |
* New (co)datatype package:
|
blanchet@61345
|
342 |
- New commands "lift_bnf" and "copy_bnf" for lifting (copying) a BNF
|
blanchet@61345
|
343 |
structure on the raw type to an abstract type defined using typedef.
|
blanchet@61345
|
344 |
- Always generate "case_transfer" theorem.
|
traytel@60920
|
345 |
|
kuncar@61370
|
346 |
* Transfer:
|
kuncar@61370
|
347 |
- new methods for interactive debugging of 'transfer' and
|
kuncar@61370
|
348 |
'transfer_prover': 'transfer_start', 'transfer_step', 'transfer_end',
|
kuncar@61370
|
349 |
'transfer_prover_start' and 'transfer_prover_end'.
|
kuncar@61370
|
350 |
|
haftmann@60868
|
351 |
* Division on integers is bootstrapped directly from division on
|
haftmann@60868
|
352 |
naturals and uses generic numeral algorithm for computations.
|
haftmann@60868
|
353 |
Slight INCOMPATIBILITY, simproc numeral_divmod replaces and generalizes
|
haftmann@60868
|
354 |
former simprocs binary_int_div and binary_int_mod
|
haftmann@60868
|
355 |
|
haftmann@60516
|
356 |
* Tightened specification of class semiring_no_zero_divisors. Slight
|
haftmann@60516
|
357 |
INCOMPATIBILITY.
|
haftmann@60516
|
358 |
|
haftmann@60688
|
359 |
* Class algebraic_semidom introduces common algebraic notions of
|
haftmann@60688
|
360 |
integral (semi)domains, particularly units. Although
|
haftmann@60517
|
361 |
logically subsumed by fields, is is not a super class of these
|
haftmann@60517
|
362 |
in order not to burden fields with notions that are trivial there.
|
haftmann@60688
|
363 |
|
haftmann@60688
|
364 |
* Class normalization_semidom specifies canonical representants
|
haftmann@60688
|
365 |
for equivalence classes of associated elements in an integral
|
haftmann@60688
|
366 |
(semi)domain. This formalizes associated elements as well.
|
haftmann@60688
|
367 |
|
haftmann@60688
|
368 |
* Abstract specification of gcd/lcm operations in classes semiring_gcd,
|
haftmann@60688
|
369 |
semiring_Gcd, semiring_Lcd. Minor INCOMPATIBILITY: facts gcd_nat.commute
|
haftmann@60688
|
370 |
and gcd_int.commute are subsumed by gcd.commute, as well as gcd_nat.assoc
|
haftmann@60688
|
371 |
and gcd_int.assoc by gcd.assoc.
|
haftmann@60517
|
372 |
|
haftmann@60429
|
373 |
* Former constants Fields.divide (_ / _) and Divides.div (_ div _)
|
haftmann@60352
|
374 |
are logically unified to Rings.divide in syntactic type class
|
haftmann@60429
|
375 |
Rings.divide, with infix syntax (_ div _). Infix syntax (_ / _)
|
haftmann@60429
|
376 |
for field division is added later as abbreviation in class Fields.inverse.
|
haftmann@60516
|
377 |
INCOMPATIBILITY, instantiations must refer to Rings.divide rather
|
haftmann@60429
|
378 |
than the former separate constants, hence infix syntax (_ / _) is usually
|
haftmann@60429
|
379 |
not available during instantiation.
|
haftmann@60352
|
380 |
|
Mathias@60397
|
381 |
* Library/Multiset:
|
Mathias@60397
|
382 |
- Renamed multiset inclusion operators:
|
Mathias@60397
|
383 |
< ~> <#
|
Mathias@60397
|
384 |
\<subset> ~> \<subset>#
|
Mathias@60397
|
385 |
<= ~> <=#
|
Mathias@60397
|
386 |
\<le> ~> \<le>#
|
Mathias@60397
|
387 |
\<subseteq> ~> \<subseteq>#
|
Mathias@60397
|
388 |
INCOMPATIBILITY.
|
Mathias@60397
|
389 |
- "'a multiset" is no longer an instance of the "order",
|
Mathias@60397
|
390 |
"ordered_ab_semigroup_add_imp_le", "ordered_cancel_comm_monoid_diff",
|
Mathias@60397
|
391 |
"semilattice_inf", and "semilattice_sup" type classes. The theorems
|
Mathias@60397
|
392 |
previously provided by these type classes (directly or indirectly)
|
Mathias@60397
|
393 |
are now available through the "subset_mset" interpretation
|
Mathias@60397
|
394 |
(e.g. add_mono ~> subset_mset.add_mono).
|
Mathias@60397
|
395 |
INCOMPATIBILITY.
|
nipkow@60497
|
396 |
- Renamed conversions:
|
nipkow@60515
|
397 |
multiset_of ~> mset
|
nipkow@60515
|
398 |
multiset_of_set ~> mset_set
|
nipkow@60497
|
399 |
set_of ~> set_mset
|
nipkow@60497
|
400 |
INCOMPATIBILITY
|
Mathias@60398
|
401 |
- Renamed lemmas:
|
Mathias@60398
|
402 |
mset_le_def ~> subseteq_mset_def
|
Mathias@60398
|
403 |
mset_less_def ~> subset_mset_def
|
Mathias@60400
|
404 |
less_eq_multiset.rep_eq ~> subseteq_mset_def
|
Mathias@60400
|
405 |
INCOMPATIBILITY
|
Mathias@60400
|
406 |
- Removed lemmas generated by lift_definition:
|
Mathias@60400
|
407 |
less_eq_multiset.abs_eq, less_eq_multiset.rsp less_eq_multiset.transfer
|
Mathias@60400
|
408 |
less_eq_multiset_def
|
Mathias@60400
|
409 |
INCOMPATIBILITY
|
wenzelm@60006
|
410 |
|
lp15@60809
|
411 |
* Multivariate_Analysis/Cauchy_Integral_Thm: Complex path integrals and Cauchy's integral theorem,
|
lp15@60809
|
412 |
ported from HOL Light
|
lp15@60809
|
413 |
|
wenzelm@60523
|
414 |
* Theory Library/Old_Recdef: discontinued obsolete 'defer_recdef'
|
wenzelm@60523
|
415 |
command. Minor INCOMPATIBILITY, use 'function' instead.
|
wenzelm@60523
|
416 |
|
wenzelm@61121
|
417 |
* Recursive function definitions ('fun', 'function', 'partial_function')
|
wenzelm@61121
|
418 |
no longer expose the low-level "_def" facts of the internal
|
wenzelm@61121
|
419 |
construction. INCOMPATIBILITY, enable option "function_defs" in the
|
wenzelm@61121
|
420 |
context for rare situations where these facts are really needed.
|
wenzelm@61121
|
421 |
|
wenzelm@61119
|
422 |
* Imperative_HOL: obsolete theory Legacy_Mrec has been removed.
|
wenzelm@61119
|
423 |
|
lammich@61178
|
424 |
* Library/Omega_Words_Fun: Infinite words modeled as functions nat => 'a.
|
lammich@61178
|
425 |
|
wenzelm@60479
|
426 |
|
wenzelm@60793
|
427 |
*** ML ***
|
wenzelm@60793
|
428 |
|
wenzelm@61268
|
429 |
* The auxiliary module Pure/display.ML has been eliminated. Its
|
wenzelm@61268
|
430 |
elementary thm print operations are now in Pure/more_thm.ML and thus
|
wenzelm@61268
|
431 |
called Thm.pretty_thm, Thm.string_of_thm etc. INCOMPATIBILITY.
|
wenzelm@61268
|
432 |
|
wenzelm@61144
|
433 |
* Simproc programming interfaces have been simplified:
|
wenzelm@61144
|
434 |
Simplifier.make_simproc and Simplifier.define_simproc supersede various
|
wenzelm@61144
|
435 |
forms of Simplifier.mk_simproc, Simplifier.simproc_global etc. Note that
|
wenzelm@61144
|
436 |
term patterns for the left-hand sides are specified with implicitly
|
wenzelm@61144
|
437 |
fixed variables, like top-level theorem statements. INCOMPATIBILITY.
|
wenzelm@61144
|
438 |
|
wenzelm@60802
|
439 |
* Instantiation rules have been re-organized as follows:
|
wenzelm@60802
|
440 |
|
wenzelm@60802
|
441 |
Thm.instantiate (*low-level instantiation with named arguments*)
|
wenzelm@60802
|
442 |
Thm.instantiate' (*version with positional arguments*)
|
wenzelm@60802
|
443 |
|
wenzelm@60802
|
444 |
Drule.infer_instantiate (*instantiation with type inference*)
|
wenzelm@60802
|
445 |
Drule.infer_instantiate' (*version with positional arguments*)
|
wenzelm@60802
|
446 |
|
wenzelm@60802
|
447 |
The LHS only requires variable specifications, instead of full terms.
|
wenzelm@60802
|
448 |
Old cterm_instantiate is superseded by infer_instantiate.
|
wenzelm@60802
|
449 |
INCOMPATIBILITY, need to re-adjust some ML names and types accordingly.
|
wenzelm@60802
|
450 |
|
wenzelm@60793
|
451 |
* Old tactic shorthands atac, rtac, etac, dtac, ftac have been
|
wenzelm@60793
|
452 |
discontinued. INCOMPATIBILITY, use regular assume_tac, resolve_tac etc.
|
wenzelm@60793
|
453 |
instead (with proper context).
|
wenzelm@60642
|
454 |
|
wenzelm@60642
|
455 |
* Thm.instantiate (and derivatives) no longer require the LHS of the
|
wenzelm@60642
|
456 |
instantiation to be certified: plain variables are given directly.
|
wenzelm@60642
|
457 |
|
wenzelm@60707
|
458 |
* Subgoal.SUBPROOF and Subgoal.FOCUS combinators use anonymous
|
wenzelm@60707
|
459 |
quasi-bound variables (like the Simplifier), instead of accidentally
|
wenzelm@60707
|
460 |
named local fixes. This has the potential to improve stability of proof
|
wenzelm@60707
|
461 |
tools, but can also cause INCOMPATIBILITY for tools that don't observe
|
wenzelm@60707
|
462 |
the proof context discipline.
|
wenzelm@60707
|
463 |
|
wenzelm@60642
|
464 |
|
wenzelm@60983
|
465 |
*** System ***
|
wenzelm@60983
|
466 |
|
wenzelm@61174
|
467 |
* Property values in etc/symbols may contain spaces, if written with the
|
wenzelm@61174
|
468 |
replacement character "␣" (Unicode point 0x2324). For example:
|
wenzelm@61174
|
469 |
|
wenzelm@61174
|
470 |
\<star> code: 0x0022c6 group: operator font: Deja␣Vu␣Sans␣Mono
|
wenzelm@61174
|
471 |
|
wenzelm@61173
|
472 |
* Command-line tool "isabelle jedit_client" allows to connect to already
|
wenzelm@61173
|
473 |
running Isabelle/jEdit process. This achieves the effect of
|
wenzelm@61173
|
474 |
single-instance applications seen on common GUI desktops.
|
wenzelm@61173
|
475 |
|
wenzelm@61216
|
476 |
* Command-line tool "isabelle update_then" expands old Isar command
|
wenzelm@61216
|
477 |
conflations:
|
wenzelm@61216
|
478 |
|
wenzelm@61216
|
479 |
hence ~> then have
|
wenzelm@61216
|
480 |
thus ~> then show
|
wenzelm@61216
|
481 |
|
wenzelm@61216
|
482 |
This syntax is more orthogonal and improves readability and
|
wenzelm@61216
|
483 |
maintainability of proofs.
|
wenzelm@61216
|
484 |
|
wenzelm@61158
|
485 |
* Poly/ML default platform architecture may be changed from 32bit to
|
wenzelm@61158
|
486 |
64bit via system option ML_system_64. A system restart (and rebuild)
|
wenzelm@61158
|
487 |
is required after change.
|
wenzelm@61158
|
488 |
|
wenzelm@61074
|
489 |
* Poly/ML 5.5.3 runs natively on x86-windows and x86_64-windows,
|
wenzelm@61074
|
490 |
which both allow larger heap space than former x86-cygwin.
|
wenzelm@60983
|
491 |
|
wenzelm@60995
|
492 |
* Java runtime environment for x86_64-windows allows to use larger heap
|
wenzelm@60995
|
493 |
space.
|
wenzelm@60995
|
494 |
|
wenzelm@61135
|
495 |
* Java runtime options are determined separately for 32bit vs. 64bit
|
wenzelm@61135
|
496 |
platforms as follows.
|
wenzelm@61135
|
497 |
|
wenzelm@61135
|
498 |
- Isabelle desktop application: platform-specific files that are
|
wenzelm@61135
|
499 |
associated with the main app bundle
|
wenzelm@61135
|
500 |
|
wenzelm@61135
|
501 |
- isabelle jedit: settings
|
wenzelm@61135
|
502 |
JEDIT_JAVA_SYSTEM_OPTIONS
|
wenzelm@61135
|
503 |
JEDIT_JAVA_OPTIONS32 vs. JEDIT_JAVA_OPTIONS64
|
wenzelm@61135
|
504 |
|
wenzelm@61135
|
505 |
- isabelle build: settings
|
wenzelm@61135
|
506 |
ISABELLE_BUILD_JAVA_OPTIONS32 vs. ISABELLE_BUILD_JAVA_OPTIONS64
|
wenzelm@61135
|
507 |
|
wenzelm@61294
|
508 |
* Bash shell function "jvmpath" has been renamed to "platform_path": it
|
wenzelm@61294
|
509 |
is relevant both for Poly/ML and JVM processes.
|
wenzelm@61294
|
510 |
|
wenzelm@60983
|
511 |
|
wenzelm@60479
|
512 |
|
wenzelm@60009
|
513 |
New in Isabelle2015 (May 2015)
|
wenzelm@60009
|
514 |
------------------------------
|
wenzelm@57695
|
515 |
|
wenzelm@57941
|
516 |
*** General ***
|
wenzelm@57941
|
517 |
|
wenzelm@59939
|
518 |
* Local theory specification commands may have a 'private' or
|
wenzelm@59990
|
519 |
'qualified' modifier to restrict name space accesses to the local scope,
|
wenzelm@59939
|
520 |
as provided by some "context begin ... end" block. For example:
|
wenzelm@59926
|
521 |
|
wenzelm@59926
|
522 |
context
|
wenzelm@59926
|
523 |
begin
|
wenzelm@59926
|
524 |
|
wenzelm@59926
|
525 |
private definition ...
|
wenzelm@59926
|
526 |
private lemma ...
|
wenzelm@59926
|
527 |
|
wenzelm@59990
|
528 |
qualified definition ...
|
wenzelm@59990
|
529 |
qualified lemma ...
|
wenzelm@59990
|
530 |
|
wenzelm@59926
|
531 |
lemma ...
|
wenzelm@59926
|
532 |
theorem ...
|
wenzelm@59926
|
533 |
|
wenzelm@59926
|
534 |
end
|
wenzelm@59926
|
535 |
|
wenzelm@59901
|
536 |
* Command 'experiment' opens an anonymous locale context with private
|
wenzelm@59901
|
537 |
naming policy.
|
wenzelm@59901
|
538 |
|
wenzelm@59951
|
539 |
* Command 'notepad' requires proper nesting of begin/end and its proof
|
wenzelm@59951
|
540 |
structure in the body: 'oops' is no longer supported here. Minor
|
wenzelm@59951
|
541 |
INCOMPATIBILITY, use 'sorry' instead.
|
wenzelm@59951
|
542 |
|
wenzelm@59951
|
543 |
* Command 'named_theorems' declares a dynamic fact within the context,
|
wenzelm@59951
|
544 |
together with an attribute to maintain the content incrementally. This
|
wenzelm@59951
|
545 |
supersedes functor Named_Thms in Isabelle/ML, but with a subtle change
|
wenzelm@59951
|
546 |
of semantics due to external visual order vs. internal reverse order.
|
wenzelm@59951
|
547 |
|
wenzelm@59951
|
548 |
* 'find_theorems': search patterns which are abstractions are
|
wenzelm@59951
|
549 |
schematically expanded before search. Search results match the naive
|
wenzelm@59951
|
550 |
expectation more closely, particularly wrt. abbreviations.
|
wenzelm@59951
|
551 |
INCOMPATIBILITY.
|
wenzelm@59648
|
552 |
|
wenzelm@59569
|
553 |
* Commands 'method_setup' and 'attribute_setup' now work within a local
|
wenzelm@59569
|
554 |
theory context.
|
wenzelm@57941
|
555 |
|
wenzelm@58928
|
556 |
* Outer syntax commands are managed authentically within the theory
|
wenzelm@59569
|
557 |
context, without implicit global state. Potential for accidental
|
wenzelm@58928
|
558 |
INCOMPATIBILITY, make sure that required theories are really imported.
|
wenzelm@58928
|
559 |
|
wenzelm@60115
|
560 |
* Historical command-line terminator ";" is no longer accepted (and
|
wenzelm@60115
|
561 |
already used differently in Isar). Minor INCOMPATIBILITY, use "isabelle
|
wenzelm@60115
|
562 |
update_semicolons" to remove obsolete semicolons from old theory
|
wenzelm@60115
|
563 |
sources.
|
wenzelm@60115
|
564 |
|
wenzelm@59951
|
565 |
* Structural composition of proof methods (meth1; meth2) in Isar
|
wenzelm@59951
|
566 |
corresponds to (tac1 THEN_ALL_NEW tac2) in ML.
|
haftmann@59105
|
567 |
|
wenzelm@60119
|
568 |
* The Eisbach proof method language allows to define new proof methods
|
wenzelm@60119
|
569 |
by combining existing ones with their usual syntax. The "match" proof
|
wenzelm@60119
|
570 |
method provides basic fact/term matching in addition to
|
wenzelm@60119
|
571 |
premise/conclusion matching through Subgoal.focus, and binds fact names
|
wenzelm@60288
|
572 |
from matches as well as term patterns within matches. The Isabelle
|
wenzelm@60288
|
573 |
documentation provides an entry "eisbach" for the Eisbach User Manual.
|
wenzelm@60288
|
574 |
Sources and various examples are in ~~/src/HOL/Eisbach/.
|
wenzelm@60119
|
575 |
|
wenzelm@57941
|
576 |
|
wenzelm@58524
|
577 |
*** Prover IDE -- Isabelle/Scala/jEdit ***
|
wenzelm@58524
|
578 |
|
wenzelm@59569
|
579 |
* Improved folding mode "isabelle" based on Isar syntax. Alternatively,
|
wenzelm@59569
|
580 |
the "sidekick" mode may be used for document structure.
|
wenzelm@59569
|
581 |
|
wenzelm@59569
|
582 |
* Extended bracket matching based on Isar language structure. System
|
wenzelm@59569
|
583 |
option jedit_structure_limit determines maximum number of lines to scan
|
wenzelm@59569
|
584 |
in the buffer.
|
wenzelm@58758
|
585 |
|
wenzelm@58540
|
586 |
* Support for BibTeX files: context menu, context-sensitive token
|
wenzelm@58540
|
587 |
marker, SideKick parser.
|
wenzelm@58524
|
588 |
|
wenzelm@58551
|
589 |
* Document antiquotation @{cite} provides formal markup, which is
|
wenzelm@60265
|
590 |
interpreted semi-formally based on .bib files that happen to be open in
|
wenzelm@60265
|
591 |
the editor (hyperlinks, completion etc.).
|
wenzelm@58551
|
592 |
|
wenzelm@58785
|
593 |
* Less waste of vertical space via negative line spacing (see Global
|
wenzelm@58785
|
594 |
Options / Text Area).
|
wenzelm@58785
|
595 |
|
wenzelm@60089
|
596 |
* Improved graphview panel with optional output of PNG or PDF, for
|
wenzelm@60273
|
597 |
display of 'thy_deps', 'class_deps' etc.
|
wenzelm@60009
|
598 |
|
wenzelm@60115
|
599 |
* The commands 'thy_deps' and 'class_deps' allow optional bounds to
|
wenzelm@60115
|
600 |
restrict the visualized hierarchy.
|
wenzelm@60093
|
601 |
|
wenzelm@60072
|
602 |
* Improved scheduling for asynchronous print commands (e.g. provers
|
wenzelm@60072
|
603 |
managed by the Sledgehammer panel) wrt. ongoing document processing.
|
wenzelm@60072
|
604 |
|
wenzelm@58524
|
605 |
|
wenzelm@59951
|
606 |
*** Document preparation ***
|
wenzelm@59951
|
607 |
|
wenzelm@59951
|
608 |
* Document markup commands 'chapter', 'section', 'subsection',
|
wenzelm@59951
|
609 |
'subsubsection', 'text', 'txt', 'text_raw' work uniformly in any
|
wenzelm@59951
|
610 |
context, even before the initial 'theory' command. Obsolete proof
|
wenzelm@59951
|
611 |
commands 'sect', 'subsect', 'subsubsect', 'txt_raw' have been
|
wenzelm@59951
|
612 |
discontinued, use 'section', 'subsection', 'subsubsection', 'text_raw'
|
wenzelm@59951
|
613 |
instead. The old 'header' command is still retained for some time, but
|
wenzelm@59951
|
614 |
should be replaced by 'chapter', 'section' etc. (using "isabelle
|
wenzelm@59951
|
615 |
update_header"). Minor INCOMPATIBILITY.
|
wenzelm@59951
|
616 |
|
wenzelm@60009
|
617 |
* Official support for "tt" style variants, via \isatt{...} or
|
wenzelm@60009
|
618 |
\begin{isabellett}...\end{isabellett}. The somewhat fragile \verb or
|
wenzelm@60009
|
619 |
verbatim environment of LaTeX is no longer used. This allows @{ML} etc.
|
wenzelm@60009
|
620 |
as argument to other macros (such as footnotes).
|
wenzelm@60009
|
621 |
|
wenzelm@60009
|
622 |
* Document antiquotation @{verbatim} prints ASCII text literally in "tt"
|
wenzelm@60009
|
623 |
style.
|
wenzelm@60009
|
624 |
|
wenzelm@60009
|
625 |
* Discontinued obsolete option "document_graph": session_graph.pdf is
|
wenzelm@60009
|
626 |
produced unconditionally for HTML browser_info and PDF-LaTeX document.
|
wenzelm@60009
|
627 |
|
wenzelm@59951
|
628 |
* Diagnostic commands and document markup commands within a proof do not
|
wenzelm@59951
|
629 |
affect the command tag for output. Thus commands like 'thm' are subject
|
wenzelm@59951
|
630 |
to proof document structure, and no longer "stick out" accidentally.
|
wenzelm@59951
|
631 |
Commands 'text' and 'txt' merely differ in the LaTeX style, not their
|
wenzelm@59951
|
632 |
tags. Potential INCOMPATIBILITY in exotic situations.
|
wenzelm@59951
|
633 |
|
wenzelm@59951
|
634 |
* System option "pretty_margin" is superseded by "thy_output_margin",
|
wenzelm@59951
|
635 |
which is also accessible via document antiquotation option "margin".
|
wenzelm@59951
|
636 |
Only the margin for document output may be changed, but not the global
|
wenzelm@59951
|
637 |
pretty printing: that is 76 for plain console output, and adapted
|
wenzelm@59951
|
638 |
dynamically in GUI front-ends. Implementations of document
|
wenzelm@59951
|
639 |
antiquotations need to observe the margin explicitly according to
|
wenzelm@59951
|
640 |
Thy_Output.string_of_margin. Minor INCOMPATIBILITY.
|
wenzelm@59951
|
641 |
|
wenzelm@60299
|
642 |
* Specification of 'document_files' in the session ROOT file is
|
wenzelm@60299
|
643 |
mandatory for document preparation. The legacy mode with implicit
|
wenzelm@60299
|
644 |
copying of the document/ directory is no longer supported. Minor
|
wenzelm@60299
|
645 |
INCOMPATIBILITY.
|
wenzelm@60299
|
646 |
|
wenzelm@59951
|
647 |
|
haftmann@58202
|
648 |
*** Pure ***
|
haftmann@58202
|
649 |
|
wenzelm@59835
|
650 |
* Proof methods with explicit instantiation ("rule_tac", "subgoal_tac"
|
wenzelm@59835
|
651 |
etc.) allow an optional context of local variables ('for' declaration):
|
wenzelm@59835
|
652 |
these variables become schematic in the instantiated theorem; this
|
wenzelm@59835
|
653 |
behaviour is analogous to 'for' in attributes "where" and "of".
|
wenzelm@59835
|
654 |
Configuration option rule_insts_schematic (default false) controls use
|
wenzelm@59835
|
655 |
of schematic variables outside the context. Minor INCOMPATIBILITY,
|
wenzelm@59835
|
656 |
declare rule_insts_schematic = true temporarily and update to use local
|
wenzelm@59835
|
657 |
variable declarations or dummy patterns instead.
|
wenzelm@59835
|
658 |
|
wenzelm@60009
|
659 |
* Explicit instantiation via attributes "where", "of", and proof methods
|
wenzelm@60009
|
660 |
"rule_tac" with derivatives like "subgoal_tac" etc. admit dummy patterns
|
wenzelm@60009
|
661 |
("_") that stand for anonymous local variables.
|
wenzelm@60009
|
662 |
|
wenzelm@59951
|
663 |
* Generated schematic variables in standard format of exported facts are
|
wenzelm@59951
|
664 |
incremented to avoid material in the proof context. Rare
|
wenzelm@59951
|
665 |
INCOMPATIBILITY, explicit instantiation sometimes needs to refer to
|
wenzelm@59951
|
666 |
different index.
|
wenzelm@59951
|
667 |
|
wenzelm@60010
|
668 |
* Lexical separation of signed and unsigned numerals: categories "num"
|
wenzelm@60010
|
669 |
and "float" are unsigned. INCOMPATIBILITY: subtle change in precedence
|
wenzelm@60010
|
670 |
of numeral signs, particularly in expressions involving infix syntax
|
wenzelm@60010
|
671 |
like "(- 1) ^ n".
|
haftmann@58410
|
672 |
|
wenzelm@58421
|
673 |
* Old inner token category "xnum" has been discontinued. Potential
|
wenzelm@58421
|
674 |
INCOMPATIBILITY for exotic syntax: may use mixfix grammar with "num"
|
wenzelm@58421
|
675 |
token category instead.
|
wenzelm@58421
|
676 |
|
haftmann@58202
|
677 |
|
blanchet@57737
|
678 |
*** HOL ***
|
blanchet@57737
|
679 |
|
blanchet@57983
|
680 |
* New (co)datatype package:
|
blanchet@58373
|
681 |
- The 'datatype_new' command has been renamed 'datatype'. The old
|
blanchet@58373
|
682 |
command of that name is now called 'old_datatype' and is provided
|
blanchet@58373
|
683 |
by "~~/src/HOL/Library/Old_Datatype.thy". See
|
blanchet@58373
|
684 |
'isabelle doc datatypes' for information on porting.
|
blanchet@58373
|
685 |
INCOMPATIBILITY.
|
blanchet@57983
|
686 |
- Renamed theorems:
|
blanchet@57983
|
687 |
disc_corec ~> corec_disc
|
blanchet@57983
|
688 |
disc_corec_iff ~> corec_disc_iff
|
blanchet@57983
|
689 |
disc_exclude ~> distinct_disc
|
blanchet@57983
|
690 |
disc_exhaust ~> exhaust_disc
|
blanchet@57983
|
691 |
disc_map_iff ~> map_disc_iff
|
blanchet@57983
|
692 |
sel_corec ~> corec_sel
|
blanchet@57983
|
693 |
sel_exhaust ~> exhaust_sel
|
blanchet@57983
|
694 |
sel_map ~> map_sel
|
blanchet@57983
|
695 |
sel_set ~> set_sel
|
blanchet@57983
|
696 |
sel_split ~> split_sel
|
blanchet@57983
|
697 |
sel_split_asm ~> split_sel_asm
|
blanchet@57983
|
698 |
strong_coinduct ~> coinduct_strong
|
blanchet@57983
|
699 |
weak_case_cong ~> case_cong_weak
|
blanchet@57983
|
700 |
INCOMPATIBILITY.
|
blanchet@58192
|
701 |
- The "no_code" option to "free_constructors", "datatype_new", and
|
blanchet@58192
|
702 |
"codatatype" has been renamed "plugins del: code".
|
blanchet@58192
|
703 |
INCOMPATIBILITY.
|
blanchet@58044
|
704 |
- The rules "set_empty" have been removed. They are easy
|
blanchet@58044
|
705 |
consequences of other set rules "by auto".
|
blanchet@58044
|
706 |
INCOMPATIBILITY.
|
blanchet@58044
|
707 |
- The rule "set_cases" is now registered with the "[cases set]"
|
blanchet@57990
|
708 |
attribute. This can influence the behavior of the "cases" proof
|
blanchet@57990
|
709 |
method when more than one case rule is applicable (e.g., an
|
blanchet@57990
|
710 |
assumption is of the form "w : set ws" and the method "cases w"
|
blanchet@57990
|
711 |
is invoked). The solution is to specify the case rule explicitly
|
blanchet@57990
|
712 |
(e.g. "cases w rule: widget.exhaust").
|
blanchet@57990
|
713 |
INCOMPATIBILITY.
|
blanchet@59675
|
714 |
- Renamed theories:
|
blanchet@59675
|
715 |
BNF_Comp ~> BNF_Composition
|
blanchet@59675
|
716 |
BNF_FP_Base ~> BNF_Fixpoint_Base
|
blanchet@59675
|
717 |
BNF_GFP ~> BNF_Greatest_Fixpoint
|
blanchet@59675
|
718 |
BNF_LFP ~> BNF_Least_Fixpoint
|
blanchet@59675
|
719 |
BNF_Constructions_on_Wellorders ~> BNF_Wellorder_Constructions
|
blanchet@59675
|
720 |
Cardinals/Constructions_on_Wellorders ~> Cardinals/Wellorder_Constructions
|
blanchet@59675
|
721 |
INCOMPATIBILITY.
|
wenzelm@60114
|
722 |
- Lifting and Transfer setup for basic HOL types sum and prod (also
|
wenzelm@60114
|
723 |
option) is now performed by the BNF package. Theories Lifting_Sum,
|
wenzelm@60114
|
724 |
Lifting_Product and Lifting_Option from Main became obsolete and
|
wenzelm@60114
|
725 |
were removed. Changed definitions of the relators rel_prod and
|
wenzelm@60114
|
726 |
rel_sum (using inductive).
|
traytel@60111
|
727 |
INCOMPATIBILITY: use rel_prod.simps and rel_sum.simps instead
|
wenzelm@60114
|
728 |
of rel_prod_def and rel_sum_def.
|
wenzelm@60114
|
729 |
Minor INCOMPATIBILITY: (rarely used by name) transfer theorem names
|
wenzelm@60114
|
730 |
changed (e.g. map_prod_transfer ~> prod.map_transfer).
|
wenzelm@60261
|
731 |
- Parametricity theorems for map functions, relators, set functions,
|
wenzelm@60261
|
732 |
constructors, case combinators, discriminators, selectors and
|
wenzelm@60261
|
733 |
(co)recursors are automatically proved and registered as transfer
|
wenzelm@60261
|
734 |
rules.
|
blanchet@57983
|
735 |
|
blanchet@57983
|
736 |
* Old datatype package:
|
blanchet@58310
|
737 |
- The old 'datatype' command has been renamed 'old_datatype', and
|
blanchet@58373
|
738 |
'rep_datatype' has been renamed 'old_rep_datatype'. They are
|
blanchet@58373
|
739 |
provided by "~~/src/HOL/Library/Old_Datatype.thy". See
|
blanchet@58310
|
740 |
'isabelle doc datatypes' for information on porting.
|
blanchet@58373
|
741 |
INCOMPATIBILITY.
|
blanchet@57983
|
742 |
- Renamed theorems:
|
blanchet@57983
|
743 |
weak_case_cong ~> case_cong_weak
|
blanchet@57983
|
744 |
INCOMPATIBILITY.
|
blanchet@58373
|
745 |
- Renamed theory:
|
blanchet@58373
|
746 |
~~/src/HOL/Datatype.thy ~> ~~/src/HOL/Library/Old_Datatype.thy
|
blanchet@58373
|
747 |
INCOMPATIBILITY.
|
blanchet@57983
|
748 |
|
blanchet@59039
|
749 |
* Nitpick:
|
wenzelm@60010
|
750 |
- Fixed soundness bug related to the strict and non-strict subset
|
blanchet@59039
|
751 |
operations.
|
blanchet@59039
|
752 |
|
blanchet@57737
|
753 |
* Sledgehammer:
|
blanchet@59511
|
754 |
- CVC4 is now included with Isabelle instead of CVC3 and run by
|
blanchet@59511
|
755 |
default.
|
blanchet@59965
|
756 |
- Z3 is now always enabled by default, now that it is fully open
|
blanchet@59965
|
757 |
source. The "z3_non_commercial" option is discontinued.
|
blanchet@57737
|
758 |
- Minimization is now always enabled by default.
|
wenzelm@60010
|
759 |
Removed sub-command:
|
blanchet@57737
|
760 |
min
|
blanchet@59967
|
761 |
- Proof reconstruction, both one-liners and Isar, has been
|
blanchet@59039
|
762 |
dramatically improved.
|
blanchet@59039
|
763 |
- Improved support for CVC4 and veriT.
|
blanchet@57737
|
764 |
|
blanchet@58062
|
765 |
* Old and new SMT modules:
|
blanchet@58067
|
766 |
- The old 'smt' method has been renamed 'old_smt' and moved to
|
wenzelm@59569
|
767 |
'src/HOL/Library/Old_SMT.thy'. It is provided for compatibility,
|
wenzelm@59569
|
768 |
until applications have been ported to use the new 'smt' method. For
|
wenzelm@59569
|
769 |
the method to work, an older version of Z3 (e.g. Z3 3.2 or 4.0) must
|
wenzelm@59569
|
770 |
be installed, and the environment variable "OLD_Z3_SOLVER" must
|
wenzelm@59569
|
771 |
point to it.
|
blanchet@58062
|
772 |
INCOMPATIBILITY.
|
blanchet@58067
|
773 |
- The 'smt2' method has been renamed 'smt'.
|
blanchet@58060
|
774 |
INCOMPATIBILITY.
|
wenzelm@59569
|
775 |
- New option 'smt_reconstruction_step_timeout' to limit the
|
wenzelm@59569
|
776 |
reconstruction time of Z3 proof steps in the new 'smt' method.
|
boehmes@59216
|
777 |
- New option 'smt_statistics' to display statistics of the new 'smt'
|
boehmes@59216
|
778 |
method, especially runtime statistics of Z3 proof reconstruction.
|
blanchet@58060
|
779 |
|
wenzelm@60261
|
780 |
* Lifting: command 'lift_definition' allows to execute lifted constants
|
wenzelm@60261
|
781 |
that have as a return type a datatype containing a subtype. This
|
wenzelm@60261
|
782 |
overcomes long-time limitations in the area of code generation and
|
wenzelm@60261
|
783 |
lifting, and avoids tedious workarounds.
|
kuncar@60258
|
784 |
|
wenzelm@60009
|
785 |
* Command and antiquotation "value" provide different evaluation slots
|
wenzelm@60009
|
786 |
(again), where the previous strategy (NBE after ML) serves as default.
|
wenzelm@60009
|
787 |
Minor INCOMPATIBILITY.
|
wenzelm@60009
|
788 |
|
wenzelm@60009
|
789 |
* Add NO_MATCH-simproc, allows to check for syntactic non-equality.
|
wenzelm@60009
|
790 |
|
wenzelm@60009
|
791 |
* field_simps: Use NO_MATCH-simproc for distribution rules, to avoid
|
wenzelm@60009
|
792 |
non-termination in case of distributing a division. With this change
|
wenzelm@60009
|
793 |
field_simps is in some cases slightly less powerful, if it fails try to
|
wenzelm@60009
|
794 |
add algebra_simps, or use divide_simps. Minor INCOMPATIBILITY.
|
wenzelm@60009
|
795 |
|
wenzelm@60009
|
796 |
* Separate class no_zero_divisors has been given up in favour of fully
|
wenzelm@60009
|
797 |
algebraic semiring_no_zero_divisors. INCOMPATIBILITY.
|
wenzelm@60009
|
798 |
|
wenzelm@60009
|
799 |
* Class linordered_semidom really requires no zero divisors.
|
wenzelm@60009
|
800 |
INCOMPATIBILITY.
|
wenzelm@60009
|
801 |
|
wenzelm@60009
|
802 |
* Classes division_ring, field and linordered_field always demand
|
wenzelm@60009
|
803 |
"inverse 0 = 0". Given up separate classes division_ring_inverse_zero,
|
wenzelm@60009
|
804 |
field_inverse_zero and linordered_field_inverse_zero. INCOMPATIBILITY.
|
wenzelm@60009
|
805 |
|
wenzelm@60009
|
806 |
* Classes cancel_ab_semigroup_add / cancel_monoid_add specify explicit
|
wenzelm@60009
|
807 |
additive inverse operation. INCOMPATIBILITY.
|
wenzelm@60009
|
808 |
|
lp15@60020
|
809 |
* Complex powers and square roots. The functions "ln" and "powr" are now
|
wenzelm@60025
|
810 |
overloaded for types real and complex, and 0 powr y = 0 by definition.
|
wenzelm@60025
|
811 |
INCOMPATIBILITY: type constraints may be necessary.
|
lp15@60020
|
812 |
|
wenzelm@60009
|
813 |
* The functions "sin" and "cos" are now defined for any type of sort
|
wenzelm@60009
|
814 |
"{real_normed_algebra_1,banach}" type, so in particular on "real" and
|
wenzelm@60009
|
815 |
"complex" uniformly. Minor INCOMPATIBILITY: type constraints may be
|
wenzelm@60009
|
816 |
needed.
|
wenzelm@60009
|
817 |
|
wenzelm@60009
|
818 |
* New library of properties of the complex transcendental functions sin,
|
wenzelm@60009
|
819 |
cos, tan, exp, Ln, Arctan, Arcsin, Arccos. Ported from HOL Light.
|
wenzelm@60009
|
820 |
|
wenzelm@60009
|
821 |
* The factorial function, "fact", now has type "nat => 'a" (of a sort
|
wenzelm@60009
|
822 |
that admits numeric types including nat, int, real and complex.
|
wenzelm@60009
|
823 |
INCOMPATIBILITY: an expression such as "fact 3 = 6" may require a type
|
wenzelm@60009
|
824 |
constraint, and the combination "real (fact k)" is likely to be
|
wenzelm@60009
|
825 |
unsatisfactory. If a type conversion is still necessary, then use
|
wenzelm@60009
|
826 |
"of_nat (fact k)" or "real_of_nat (fact k)".
|
wenzelm@60009
|
827 |
|
wenzelm@60009
|
828 |
* Removed functions "natfloor" and "natceiling", use "nat o floor" and
|
wenzelm@60009
|
829 |
"nat o ceiling" instead. A few of the lemmas have been retained and
|
wenzelm@60009
|
830 |
adapted: in their names "natfloor"/"natceiling" has been replaced by
|
wenzelm@60009
|
831 |
"nat_floor"/"nat_ceiling".
|
wenzelm@60009
|
832 |
|
wenzelm@60009
|
833 |
* Qualified some duplicated fact names required for boostrapping the
|
wenzelm@60009
|
834 |
type class hierarchy:
|
wenzelm@60009
|
835 |
ab_add_uminus_conv_diff ~> diff_conv_add_uminus
|
wenzelm@60009
|
836 |
field_inverse_zero ~> inverse_zero
|
wenzelm@60009
|
837 |
field_divide_inverse ~> divide_inverse
|
wenzelm@60009
|
838 |
field_inverse ~> left_inverse
|
wenzelm@60009
|
839 |
Minor INCOMPATIBILITY.
|
wenzelm@60009
|
840 |
|
wenzelm@60009
|
841 |
* Eliminated fact duplicates:
|
wenzelm@60009
|
842 |
mult_less_imp_less_right ~> mult_right_less_imp_less
|
wenzelm@60009
|
843 |
mult_less_imp_less_left ~> mult_left_less_imp_less
|
wenzelm@60009
|
844 |
Minor INCOMPATIBILITY.
|
wenzelm@60009
|
845 |
|
wenzelm@60009
|
846 |
* Fact consolidation: even_less_0_iff is subsumed by
|
wenzelm@60009
|
847 |
double_add_less_zero_iff_single_add_less_zero (simp by default anyway).
|
wenzelm@60009
|
848 |
|
wenzelm@60009
|
849 |
* Generalized and consolidated some theorems concerning divsibility:
|
wenzelm@60009
|
850 |
dvd_reduce ~> dvd_add_triv_right_iff
|
wenzelm@60009
|
851 |
dvd_plus_eq_right ~> dvd_add_right_iff
|
wenzelm@60009
|
852 |
dvd_plus_eq_left ~> dvd_add_left_iff
|
wenzelm@60009
|
853 |
Minor INCOMPATIBILITY.
|
wenzelm@60009
|
854 |
|
wenzelm@60009
|
855 |
* "even" and "odd" are mere abbreviations for "2 dvd _" and "~ 2 dvd _"
|
wenzelm@60009
|
856 |
and part of theory Main.
|
wenzelm@60009
|
857 |
even_def ~> even_iff_mod_2_eq_zero
|
wenzelm@60009
|
858 |
INCOMPATIBILITY.
|
wenzelm@60009
|
859 |
|
wenzelm@60009
|
860 |
* Lemma name consolidation: divide_Numeral1 ~> divide_numeral_1. Minor
|
wenzelm@60009
|
861 |
INCOMPATIBILITY.
|
wenzelm@60009
|
862 |
|
wenzelm@60009
|
863 |
* Bootstrap of listsum as special case of abstract product over lists.
|
wenzelm@60009
|
864 |
Fact rename:
|
wenzelm@60009
|
865 |
listsum_def ~> listsum.eq_foldr
|
wenzelm@60009
|
866 |
INCOMPATIBILITY.
|
wenzelm@60009
|
867 |
|
wenzelm@60009
|
868 |
* Product over lists via constant "listprod".
|
wenzelm@60009
|
869 |
|
wenzelm@60009
|
870 |
* Theory List: renamed drop_Suc_conv_tl and nth_drop' to
|
wenzelm@60009
|
871 |
Cons_nth_drop_Suc.
|
nipkow@58247
|
872 |
|
Andreas@58626
|
873 |
* New infrastructure for compiling, running, evaluating and testing
|
wenzelm@59569
|
874 |
generated code in target languages in HOL/Library/Code_Test. See
|
wenzelm@59569
|
875 |
HOL/Codegenerator_Test/Code_Test* for examples.
|
wenzelm@58008
|
876 |
|
wenzelm@60009
|
877 |
* Library/Multiset:
|
blanchet@59813
|
878 |
- Introduced "replicate_mset" operation.
|
blanchet@59813
|
879 |
- Introduced alternative characterizations of the multiset ordering in
|
blanchet@59813
|
880 |
"Library/Multiset_Order".
|
blanchet@59958
|
881 |
- Renamed multiset ordering:
|
blanchet@59958
|
882 |
<# ~> #<#
|
blanchet@59958
|
883 |
<=# ~> #<=#
|
blanchet@59958
|
884 |
\<subset># ~> #\<subset>#
|
blanchet@59958
|
885 |
\<subseteq># ~> #\<subseteq>#
|
blanchet@59958
|
886 |
INCOMPATIBILITY.
|
blanchet@59986
|
887 |
- Introduced abbreviations for ill-named multiset operations:
|
blanchet@59986
|
888 |
<#, \<subset># abbreviate < (strict subset)
|
blanchet@59986
|
889 |
<=#, \<le>#, \<subseteq># abbreviate <= (subset or equal)
|
blanchet@59986
|
890 |
INCOMPATIBILITY.
|
blanchet@59813
|
891 |
- Renamed
|
blanchet@59813
|
892 |
in_multiset_of ~> in_multiset_in_set
|
nipkow@59998
|
893 |
Multiset.fold ~> fold_mset
|
nipkow@59998
|
894 |
Multiset.filter ~> filter_mset
|
blanchet@59813
|
895 |
INCOMPATIBILITY.
|
nipkow@59949
|
896 |
- Removed mcard, is equal to size.
|
blanchet@59813
|
897 |
- Added attributes:
|
blanchet@59813
|
898 |
image_mset.id [simp]
|
blanchet@59813
|
899 |
image_mset_id [simp]
|
blanchet@59813
|
900 |
elem_multiset_of_set [simp, intro]
|
blanchet@59813
|
901 |
comp_fun_commute_plus_mset [simp]
|
blanchet@59813
|
902 |
comp_fun_commute.fold_mset_insert [OF comp_fun_commute_plus_mset, simp]
|
blanchet@59813
|
903 |
in_mset_fold_plus_iff [iff]
|
blanchet@59813
|
904 |
set_of_Union_mset [simp]
|
blanchet@59813
|
905 |
in_Union_mset_iff [iff]
|
blanchet@59813
|
906 |
INCOMPATIBILITY.
|
blanchet@59813
|
907 |
|
wenzelm@60009
|
908 |
* Library/Sum_of_Squares: simplified and improved "sos" method. Always
|
wenzelm@60009
|
909 |
use local CSDP executable, which is much faster than the NEOS server.
|
wenzelm@60009
|
910 |
The "sos_cert" functionality is invoked as "sos" with additional
|
wenzelm@60009
|
911 |
argument. Minor INCOMPATIBILITY.
|
wenzelm@60009
|
912 |
|
wenzelm@60009
|
913 |
* HOL-Decision_Procs: New counterexample generator quickcheck
|
wenzelm@60009
|
914 |
[approximation] for inequalities of transcendental functions. Uses
|
wenzelm@60009
|
915 |
hardware floating point arithmetic to randomly discover potential
|
wenzelm@60010
|
916 |
counterexamples. Counterexamples are certified with the "approximation"
|
wenzelm@60009
|
917 |
method. See HOL/Decision_Procs/ex/Approximation_Quickcheck_Ex.thy for
|
wenzelm@60009
|
918 |
examples.
|
immler@58990
|
919 |
|
hoelzl@59354
|
920 |
* HOL-Probability: Reworked measurability prover
|
wenzelm@60010
|
921 |
- applies destructor rules repeatedly
|
hoelzl@59354
|
922 |
- removed application splitting (replaced by destructor rule)
|
wenzelm@59569
|
923 |
- added congruence rules to rewrite measure spaces under the sets
|
wenzelm@59569
|
924 |
projection
|
wenzelm@59569
|
925 |
|
wenzelm@60009
|
926 |
* New proof method "rewrite" (in theory ~~/src/HOL/Library/Rewrite) for
|
wenzelm@60009
|
927 |
single-step rewriting with subterm selection based on patterns.
|
wenzelm@60009
|
928 |
|
wenzelm@58630
|
929 |
|
blanchet@58066
|
930 |
*** ML ***
|
blanchet@58066
|
931 |
|
wenzelm@60009
|
932 |
* Subtle change of name space policy: undeclared entries are now
|
wenzelm@60009
|
933 |
considered inaccessible, instead of accessible via the fully-qualified
|
wenzelm@60009
|
934 |
internal name. This mainly affects Name_Space.intern (and derivatives),
|
wenzelm@60009
|
935 |
which may produce an unexpected Long_Name.hidden prefix. Note that
|
wenzelm@60010
|
936 |
contemporary applications use the strict Name_Space.check (and
|
wenzelm@60009
|
937 |
derivatives) instead, which is not affected by the change. Potential
|
wenzelm@60009
|
938 |
INCOMPATIBILITY in rare applications of Name_Space.intern.
|
wenzelm@59951
|
939 |
|
wenzelm@60094
|
940 |
* Subtle change of error semantics of Toplevel.proof_of: regular user
|
wenzelm@60094
|
941 |
ERROR instead of internal Toplevel.UNDEF.
|
wenzelm@60094
|
942 |
|
wenzelm@59951
|
943 |
* Basic combinators map, fold, fold_map, split_list, apply are available
|
wenzelm@59951
|
944 |
as parameterized antiquotations, e.g. @{map 4} for lists of quadruples.
|
wenzelm@59951
|
945 |
|
wenzelm@59951
|
946 |
* Renamed "pairself" to "apply2", in accordance to @{apply 2}.
|
wenzelm@59951
|
947 |
INCOMPATIBILITY.
|
wenzelm@59951
|
948 |
|
wenzelm@59951
|
949 |
* Former combinators NAMED_CRITICAL and CRITICAL for central critical
|
wenzelm@59951
|
950 |
sections have been discontinued, in favour of the more elementary
|
wenzelm@59951
|
951 |
Multithreading.synchronized and its high-level derivative
|
wenzelm@59951
|
952 |
Synchronized.var (which is usually sufficient in applications). Subtle
|
wenzelm@59951
|
953 |
INCOMPATIBILITY: synchronized access needs to be atomic and cannot be
|
wenzelm@59951
|
954 |
nested.
|
wenzelm@59951
|
955 |
|
wenzelm@60009
|
956 |
* Synchronized.value (ML) is actually synchronized (as in Scala): subtle
|
wenzelm@60009
|
957 |
change of semantics with minimal potential for INCOMPATIBILITY.
|
wenzelm@59899
|
958 |
|
wenzelm@59621
|
959 |
* The main operations to certify logical entities are Thm.ctyp_of and
|
wenzelm@59621
|
960 |
Thm.cterm_of with a local context; old-style global theory variants are
|
wenzelm@59621
|
961 |
available as Thm.global_ctyp_of and Thm.global_cterm_of.
|
wenzelm@59621
|
962 |
INCOMPATIBILITY.
|
wenzelm@59621
|
963 |
|
wenzelm@59582
|
964 |
* Elementary operations in module Thm are no longer pervasive.
|
wenzelm@59582
|
965 |
INCOMPATIBILITY, need to use qualified Thm.prop_of, Thm.cterm_of,
|
wenzelm@59582
|
966 |
Thm.term_of etc.
|
wenzelm@59582
|
967 |
|
wenzelm@58963
|
968 |
* Proper context for various elementary tactics: assume_tac,
|
wenzelm@59498
|
969 |
resolve_tac, eresolve_tac, dresolve_tac, forward_tac, match_tac,
|
wenzelm@59498
|
970 |
compose_tac, Splitter.split_tac etc. INCOMPATIBILITY.
|
wenzelm@58956
|
971 |
|
blanchet@58066
|
972 |
* Tactical PARALLEL_ALLGOALS is the most common way to refer to
|
blanchet@58066
|
973 |
PARALLEL_GOALS.
|
blanchet@58066
|
974 |
|
wenzelm@59564
|
975 |
* Goal.prove_multi is superseded by the fully general Goal.prove_common,
|
wenzelm@59564
|
976 |
which also allows to specify a fork priority.
|
wenzelm@59564
|
977 |
|
wenzelm@59936
|
978 |
* Antiquotation @{command_spec "COMMAND"} is superseded by
|
wenzelm@59936
|
979 |
@{command_keyword COMMAND} (usually without quotes and with PIDE
|
wenzelm@59936
|
980 |
markup). Minor INCOMPATIBILITY.
|
wenzelm@59936
|
981 |
|
wenzelm@60009
|
982 |
* Cartouches within ML sources are turned into values of type
|
wenzelm@60009
|
983 |
Input.source (with formal position information).
|
wenzelm@60009
|
984 |
|
blanchet@58066
|
985 |
|
wenzelm@58610
|
986 |
*** System ***
|
wenzelm@58610
|
987 |
|
wenzelm@59951
|
988 |
* The Isabelle tool "update_cartouches" changes theory files to use
|
wenzelm@59951
|
989 |
cartouches instead of old-style {* verbatim *} or `alt_string` tokens.
|
wenzelm@59951
|
990 |
|
wenzelm@60106
|
991 |
* The Isabelle tool "build" provides new options -X, -k, -x.
|
wenzelm@59951
|
992 |
|
wenzelm@59951
|
993 |
* Discontinued old-fashioned "codegen" tool. Code generation can always
|
wenzelm@59951
|
994 |
be externally triggered using an appropriate ROOT file plus a
|
wenzelm@59951
|
995 |
corresponding theory. Parametrization is possible using environment
|
wenzelm@59951
|
996 |
variables, or ML snippets in the most extreme cases. Minor
|
wenzelm@59951
|
997 |
INCOMPATIBILITY.
|
wenzelm@58842
|
998 |
|
wenzelm@59200
|
999 |
* JVM system property "isabelle.threads" determines size of Scala thread
|
wenzelm@59200
|
1000 |
pool, like Isabelle system option "threads" for ML.
|
wenzelm@59200
|
1001 |
|
wenzelm@59201
|
1002 |
* JVM system property "isabelle.laf" determines the default Swing
|
wenzelm@59201
|
1003 |
look-and-feel, via internal class name or symbolic name as in the jEdit
|
wenzelm@59201
|
1004 |
menu Global Options / Appearance.
|
wenzelm@59201
|
1005 |
|
wenzelm@59951
|
1006 |
* Support for Proof General and Isar TTY loop has been discontinued.
|
wenzelm@60009
|
1007 |
Minor INCOMPATIBILITY, use standard PIDE infrastructure instead.
|
wenzelm@59891
|
1008 |
|
wenzelm@58610
|
1009 |
|
wenzelm@57695
|
1010 |
|
wenzelm@57452
|
1011 |
New in Isabelle2014 (August 2014)
|
wenzelm@57452
|
1012 |
---------------------------------
|
wenzelm@54055
|
1013 |
|
wenzelm@54702
|
1014 |
*** General ***
|
wenzelm@54702
|
1015 |
|
wenzelm@57452
|
1016 |
* Support for official Standard ML within the Isabelle context.
|
wenzelm@57452
|
1017 |
Command 'SML_file' reads and evaluates the given Standard ML file.
|
wenzelm@57452
|
1018 |
Toplevel bindings are stored within the theory context; the initial
|
wenzelm@57452
|
1019 |
environment is restricted to the Standard ML implementation of
|
wenzelm@57452
|
1020 |
Poly/ML, without the add-ons of Isabelle/ML. Commands 'SML_import'
|
wenzelm@57452
|
1021 |
and 'SML_export' allow to exchange toplevel bindings between the two
|
wenzelm@57452
|
1022 |
separate environments. See also ~~/src/Tools/SML/Examples.thy for
|
wenzelm@57452
|
1023 |
some examples.
|
wenzelm@56499
|
1024 |
|
wenzelm@57504
|
1025 |
* Standard tactics and proof methods such as "clarsimp", "auto" and
|
wenzelm@57504
|
1026 |
"safe" now preserve equality hypotheses "x = expr" where x is a free
|
wenzelm@57504
|
1027 |
variable. Locale assumptions and chained facts containing "x"
|
wenzelm@57504
|
1028 |
continue to be useful. The new method "hypsubst_thin" and the
|
wenzelm@57504
|
1029 |
configuration option "hypsubst_thin" (within the attribute name space)
|
wenzelm@57504
|
1030 |
restore the previous behavior. INCOMPATIBILITY, especially where
|
wenzelm@57504
|
1031 |
induction is done after these methods or when the names of free and
|
wenzelm@57504
|
1032 |
bound variables clash. As first approximation, old proofs may be
|
wenzelm@57504
|
1033 |
repaired by "using [[hypsubst_thin = true]]" in the critical spot.
|
wenzelm@57504
|
1034 |
|
wenzelm@56232
|
1035 |
* More static checking of proof methods, which allows the system to
|
wenzelm@56232
|
1036 |
form a closure over the concrete syntax. Method arguments should be
|
wenzelm@56232
|
1037 |
processed in the original proof context as far as possible, before
|
wenzelm@56232
|
1038 |
operating on the goal state. In any case, the standard discipline for
|
wenzelm@56232
|
1039 |
subgoal-addressing needs to be observed: no subgoals or a subgoal
|
wenzelm@56232
|
1040 |
number that is out of range produces an empty result sequence, not an
|
wenzelm@56232
|
1041 |
exception. Potential INCOMPATIBILITY for non-conformant tactical
|
wenzelm@56232
|
1042 |
proof tools.
|
wenzelm@56232
|
1043 |
|
wenzelm@57452
|
1044 |
* Lexical syntax (inner and outer) supports text cartouches with
|
wenzelm@57452
|
1045 |
arbitrary nesting, and without escapes of quotes etc. The Prover IDE
|
wenzelm@57452
|
1046 |
supports input via ` (backquote).
|
wenzelm@57452
|
1047 |
|
wenzelm@57452
|
1048 |
* The outer syntax categories "text" (for formal comments and document
|
wenzelm@57452
|
1049 |
markup commands) and "altstring" (for literal fact references) allow
|
wenzelm@57452
|
1050 |
cartouches as well, in addition to the traditional mix of quotations.
|
wenzelm@57452
|
1051 |
|
wenzelm@57452
|
1052 |
* Syntax of document antiquotation @{rail} now uses \<newline> instead
|
wenzelm@57452
|
1053 |
of "\\", to avoid the optical illusion of escaped backslash within
|
wenzelm@57491
|
1054 |
string token. General renovation of its syntax using text cartouches.
|
wenzelm@57452
|
1055 |
Minor INCOMPATIBILITY.
|
wenzelm@57452
|
1056 |
|
wenzelm@57452
|
1057 |
* Discontinued legacy_isub_isup, which was a temporary workaround for
|
wenzelm@57452
|
1058 |
Isabelle/ML in Isabelle2013-1. The prover process no longer accepts
|
wenzelm@57452
|
1059 |
old identifier syntax with \<^isub> or \<^isup>. Potential
|
wenzelm@57452
|
1060 |
INCOMPATIBILITY.
|
wenzelm@57452
|
1061 |
|
wenzelm@57452
|
1062 |
* Document antiquotation @{url} produces markup for the given URL,
|
wenzelm@57452
|
1063 |
which results in an active hyperlink within the text.
|
wenzelm@57452
|
1064 |
|
wenzelm@57452
|
1065 |
* Document antiquotation @{file_unchecked} is like @{file}, but does
|
wenzelm@57452
|
1066 |
not check existence within the file-system.
|
wenzelm@57452
|
1067 |
|
wenzelm@57452
|
1068 |
* Updated and extended manuals: codegen, datatypes, implementation,
|
wenzelm@57452
|
1069 |
isar-ref, jedit, system.
|
wenzelm@57423
|
1070 |
|
wenzelm@54702
|
1071 |
|
wenzelm@54533
|
1072 |
*** Prover IDE -- Isabelle/Scala/jEdit ***
|
wenzelm@54533
|
1073 |
|
wenzelm@57650
|
1074 |
* Improved Document panel: simplified interaction where every single
|
wenzelm@57452
|
1075 |
mouse click (re)opens document via desktop environment or as jEdit
|
wenzelm@57452
|
1076 |
buffer.
|
wenzelm@57452
|
1077 |
|
wenzelm@57452
|
1078 |
* Support for Navigator plugin (with toolbar buttons), with connection
|
wenzelm@57452
|
1079 |
to PIDE hyperlinks.
|
wenzelm@57452
|
1080 |
|
wenzelm@57452
|
1081 |
* Auxiliary files ('ML_file' etc.) are managed by the Prover IDE.
|
wenzelm@57452
|
1082 |
Open text buffers take precedence over copies within the file-system.
|
wenzelm@57452
|
1083 |
|
wenzelm@57452
|
1084 |
* Improved support for Isabelle/ML, with jEdit mode "isabelle-ml" for
|
wenzelm@57452
|
1085 |
auxiliary ML files.
|
wenzelm@57423
|
1086 |
|
wenzelm@57423
|
1087 |
* Improved syntactic and semantic completion mechanism, with simple
|
wenzelm@57423
|
1088 |
templates, completion language context, name-space completion,
|
wenzelm@57423
|
1089 |
file-name completion, spell-checker completion.
|
wenzelm@57423
|
1090 |
|
wenzelm@57423
|
1091 |
* Refined GUI popup for completion: more robust key/mouse event
|
wenzelm@57423
|
1092 |
handling and propagation to enclosing text area -- avoid loosing
|
wenzelm@57423
|
1093 |
keystrokes with slow / remote graphics displays.
|
wenzelm@57423
|
1094 |
|
wenzelm@57833
|
1095 |
* Completion popup supports both ENTER and TAB (default) to select an
|
wenzelm@57833
|
1096 |
item, depending on Isabelle options.
|
wenzelm@57833
|
1097 |
|
wenzelm@57423
|
1098 |
* Refined insertion of completion items wrt. jEdit text: multiple
|
wenzelm@57423
|
1099 |
selections, rectangular selections, rectangular selection as "tall
|
wenzelm@57423
|
1100 |
caret".
|
wenzelm@56342
|
1101 |
|
wenzelm@56580
|
1102 |
* Integrated spell-checker for document text, comments etc. with
|
wenzelm@57423
|
1103 |
completion popup and context-menu.
|
wenzelm@56554
|
1104 |
|
wenzelm@56879
|
1105 |
* More general "Query" panel supersedes "Find" panel, with GUI access
|
wenzelm@56879
|
1106 |
to commands 'find_theorems' and 'find_consts', as well as print
|
wenzelm@56879
|
1107 |
operations for the context. Minor incompatibility in keyboard
|
wenzelm@56879
|
1108 |
shortcuts etc.: replace action isabelle-find by isabelle-query.
|
wenzelm@56761
|
1109 |
|
wenzelm@56901
|
1110 |
* Search field for all output panels ("Output", "Query", "Info" etc.)
|
wenzelm@56901
|
1111 |
to highlight text via regular expression.
|
wenzelm@56901
|
1112 |
|
wenzelm@54881
|
1113 |
* Option "jedit_print_mode" (see also "Plugin Options / Isabelle /
|
wenzelm@54881
|
1114 |
General") allows to specify additional print modes for the prover
|
wenzelm@54881
|
1115 |
process, without requiring old-fashioned command-line invocation of
|
wenzelm@54881
|
1116 |
"isabelle jedit -m MODE".
|
wenzelm@54881
|
1117 |
|
wenzelm@56505
|
1118 |
* More support for remote files (e.g. http) using standard Java
|
wenzelm@56505
|
1119 |
networking operations instead of jEdit virtual file-systems.
|
wenzelm@56505
|
1120 |
|
wenzelm@57822
|
1121 |
* Empty editors buffers that are no longer required (e.g.\ via theory
|
wenzelm@57822
|
1122 |
imports) are automatically removed from the document model.
|
wenzelm@57822
|
1123 |
|
wenzelm@57869
|
1124 |
* Improved monitor panel.
|
wenzelm@57869
|
1125 |
|
wenzelm@56838
|
1126 |
* Improved Console/Scala plugin: more uniform scala.Console output,
|
wenzelm@56838
|
1127 |
more robust treatment of threads and interrupts.
|
wenzelm@56838
|
1128 |
|
wenzelm@56939
|
1129 |
* Improved management of dockable windows: clarified keyboard focus
|
wenzelm@56939
|
1130 |
and window placement wrt. main editor view; optional menu item to
|
wenzelm@56939
|
1131 |
"Detach" a copy where this makes sense.
|
wenzelm@56939
|
1132 |
|
wenzelm@57452
|
1133 |
* New Simplifier Trace panel provides an interactive view of the
|
wenzelm@57591
|
1134 |
simplification process, enabled by the "simp_trace_new" attribute
|
wenzelm@57452
|
1135 |
within the context.
|
wenzelm@57452
|
1136 |
|
wenzelm@57452
|
1137 |
|
wenzelm@55001
|
1138 |
*** Pure ***
|
wenzelm@55001
|
1139 |
|
wenzelm@57504
|
1140 |
* Low-level type-class commands 'classes', 'classrel', 'arities' have
|
wenzelm@57504
|
1141 |
been discontinued to avoid the danger of non-trivial axiomatization
|
wenzelm@57504
|
1142 |
that is not immediately visible. INCOMPATIBILITY, use regular
|
wenzelm@57504
|
1143 |
'instance' command with proof. The required OFCLASS(...) theorem
|
wenzelm@57504
|
1144 |
might be postulated via 'axiomatization' beforehand, or the proof
|
wenzelm@57504
|
1145 |
finished trivially if the underlying class definition is made vacuous
|
wenzelm@57504
|
1146 |
(without any assumptions). See also Isabelle/ML operations
|
wenzelm@57504
|
1147 |
Axclass.class_axiomatization, Axclass.classrel_axiomatization,
|
wenzelm@57504
|
1148 |
Axclass.arity_axiomatization.
|
wenzelm@57504
|
1149 |
|
wenzelm@56245
|
1150 |
* Basic constants of Pure use more conventional names and are always
|
wenzelm@56245
|
1151 |
qualified. Rare INCOMPATIBILITY, but with potentially serious
|
wenzelm@56245
|
1152 |
consequences, notably for tools in Isabelle/ML. The following
|
wenzelm@56245
|
1153 |
renaming needs to be applied:
|
wenzelm@56245
|
1154 |
|
wenzelm@56245
|
1155 |
== ~> Pure.eq
|
wenzelm@56245
|
1156 |
==> ~> Pure.imp
|
wenzelm@56245
|
1157 |
all ~> Pure.all
|
wenzelm@56245
|
1158 |
TYPE ~> Pure.type
|
wenzelm@56245
|
1159 |
dummy_pattern ~> Pure.dummy_pattern
|
wenzelm@56245
|
1160 |
|
wenzelm@56245
|
1161 |
Systematic porting works by using the following theory setup on a
|
wenzelm@56245
|
1162 |
*previous* Isabelle version to introduce the new name accesses for the
|
wenzelm@56245
|
1163 |
old constants:
|
wenzelm@56245
|
1164 |
|
wenzelm@56245
|
1165 |
setup {*
|
wenzelm@56245
|
1166 |
fn thy => thy
|
wenzelm@56245
|
1167 |
|> Sign.root_path
|
wenzelm@56245
|
1168 |
|> Sign.const_alias (Binding.qualify true "Pure" @{binding eq}) "=="
|
wenzelm@56245
|
1169 |
|> Sign.const_alias (Binding.qualify true "Pure" @{binding imp}) "==>"
|
wenzelm@56245
|
1170 |
|> Sign.const_alias (Binding.qualify true "Pure" @{binding all}) "all"
|
wenzelm@56245
|
1171 |
|> Sign.restore_naming thy
|
wenzelm@56245
|
1172 |
*}
|
wenzelm@56245
|
1173 |
|
wenzelm@56245
|
1174 |
Thus ML antiquotations like @{const_name Pure.eq} may be used already.
|
wenzelm@56245
|
1175 |
Later the application is moved to the current Isabelle version, and
|
wenzelm@56245
|
1176 |
the auxiliary aliases are deleted.
|
wenzelm@56245
|
1177 |
|
wenzelm@55143
|
1178 |
* Attributes "where" and "of" allow an optional context of local
|
wenzelm@55143
|
1179 |
variables ('for' declaration): these variables become schematic in the
|
wenzelm@55143
|
1180 |
instantiated theorem.
|
wenzelm@55143
|
1181 |
|
wenzelm@55152
|
1182 |
* Obsolete attribute "standard" has been discontinued (legacy since
|
wenzelm@55152
|
1183 |
Isabelle2012). Potential INCOMPATIBILITY, use explicit 'for' context
|
wenzelm@55152
|
1184 |
where instantiations with schematic variables are intended (for
|
wenzelm@55152
|
1185 |
declaration commands like 'lemmas' or attributes like "of"). The
|
wenzelm@55152
|
1186 |
following temporary definition may help to port old applications:
|
wenzelm@55152
|
1187 |
|
wenzelm@55152
|
1188 |
attribute_setup standard =
|
wenzelm@55152
|
1189 |
"Scan.succeed (Thm.rule_attribute (K Drule.export_without_context))"
|
wenzelm@55152
|
1190 |
|
wenzelm@55001
|
1191 |
* More thorough check of proof context for goal statements and
|
wenzelm@55006
|
1192 |
attributed fact expressions (concerning background theory, declared
|
wenzelm@55006
|
1193 |
hyps). Potential INCOMPATIBILITY, tools need to observe standard
|
wenzelm@55006
|
1194 |
context discipline. See also Assumption.add_assumes and the more
|
wenzelm@55006
|
1195 |
primitive Thm.assume_hyps.
|
wenzelm@55001
|
1196 |
|
wenzelm@55108
|
1197 |
* Inner syntax token language allows regular quoted strings "..."
|
wenzelm@55108
|
1198 |
(only makes sense in practice, if outer syntax is delimited
|
wenzelm@57452
|
1199 |
differently, e.g. via cartouches).
|
wenzelm@57452
|
1200 |
|
wenzelm@57504
|
1201 |
* Command 'print_term_bindings' supersedes 'print_binds' for clarity,
|
wenzelm@57504
|
1202 |
but the latter is retained some time as Proof General legacy.
|
wenzelm@57504
|
1203 |
|
wenzelm@57452
|
1204 |
* Code generator preprocessor: explicit control of simp tracing on a
|
wenzelm@57452
|
1205 |
per-constant basis. See attribute "code_preproc".
|
haftmann@57430
|
1206 |
|
wenzelm@55001
|
1207 |
|
haftmann@54227
|
1208 |
*** HOL ***
|
haftmann@54227
|
1209 |
|
wenzelm@57504
|
1210 |
* Code generator: enforce case of identifiers only for strict target
|
wenzelm@57504
|
1211 |
language requirements. INCOMPATIBILITY.
|
wenzelm@57504
|
1212 |
|
wenzelm@57504
|
1213 |
* Code generator: explicit proof contexts in many ML interfaces.
|
wenzelm@57504
|
1214 |
INCOMPATIBILITY.
|
wenzelm@57504
|
1215 |
|
wenzelm@57504
|
1216 |
* Code generator: minimize exported identifiers by default. Minor
|
wenzelm@57504
|
1217 |
INCOMPATIBILITY.
|
wenzelm@57504
|
1218 |
|
wenzelm@57504
|
1219 |
* Code generation for SML and OCaml: dropped arcane "no_signatures"
|
wenzelm@57504
|
1220 |
option. Minor INCOMPATIBILITY.
|
wenzelm@57504
|
1221 |
|
wenzelm@57504
|
1222 |
* "declare [[code abort: ...]]" replaces "code_abort ...".
|
wenzelm@57504
|
1223 |
INCOMPATIBILITY.
|
wenzelm@57504
|
1224 |
|
wenzelm@57504
|
1225 |
* "declare [[code drop: ...]]" drops all code equations associated
|
wenzelm@57504
|
1226 |
with the given constants.
|
wenzelm@57504
|
1227 |
|
wenzelm@57504
|
1228 |
* Code generations are provided for make, fields, extend and truncate
|
wenzelm@57504
|
1229 |
operations on records.
|
haftmann@57437
|
1230 |
|
wenzelm@57452
|
1231 |
* Command and antiquotation "value" are now hardcoded against nbe and
|
wenzelm@57452
|
1232 |
ML. Minor INCOMPATIBILITY.
|
wenzelm@57452
|
1233 |
|
wenzelm@57504
|
1234 |
* Renamed command 'enriched_type' to 'functor'. INCOMPATIBILITY.
|
wenzelm@57504
|
1235 |
|
wenzelm@57504
|
1236 |
* The symbol "\<newline>" may be used within char or string literals
|
wenzelm@57504
|
1237 |
to represent (Char Nibble0 NibbleA), i.e. ASCII newline.
|
wenzelm@57504
|
1238 |
|
wenzelm@57504
|
1239 |
* Qualified String.implode and String.explode. INCOMPATIBILITY.
|
haftmann@56923
|
1240 |
|
wenzelm@57452
|
1241 |
* Simplifier: Enhanced solver of preconditions of rewrite rules can
|
wenzelm@57452
|
1242 |
now deal with conjunctions. For help with converting proofs, the old
|
wenzelm@57452
|
1243 |
behaviour of the simplifier can be restored like this: declare/using
|
wenzelm@57452
|
1244 |
[[simp_legacy_precond]]. This configuration option will disappear
|
wenzelm@57452
|
1245 |
again in the future. INCOMPATIBILITY.
|
nipkow@56073
|
1246 |
|
wenzelm@55139
|
1247 |
* Simproc "finite_Collect" is no longer enabled by default, due to
|
wenzelm@55139
|
1248 |
spurious crashes and other surprises. Potential INCOMPATIBILITY.
|
wenzelm@55139
|
1249 |
|
wenzelm@57452
|
1250 |
* Moved new (co)datatype package and its dependencies from session
|
wenzelm@57452
|
1251 |
"HOL-BNF" to "HOL". The commands 'bnf', 'wrap_free_constructors',
|
wenzelm@57452
|
1252 |
'datatype_new', 'codatatype', 'primcorec', 'primcorecursive' are now
|
wenzelm@57452
|
1253 |
part of theory "Main".
|
wenzelm@57452
|
1254 |
|
blanchet@55098
|
1255 |
Theory renamings:
|
blanchet@55098
|
1256 |
FunDef.thy ~> Fun_Def.thy (and Fun_Def_Base.thy)
|
blanchet@55098
|
1257 |
Library/Wfrec.thy ~> Wfrec.thy
|
blanchet@55098
|
1258 |
Library/Zorn.thy ~> Zorn.thy
|
blanchet@55098
|
1259 |
Cardinals/Order_Relation.thy ~> Order_Relation.thy
|
blanchet@55098
|
1260 |
Library/Order_Union.thy ~> Cardinals/Order_Union.thy
|
blanchet@55098
|
1261 |
Cardinals/Cardinal_Arithmetic_Base.thy ~> BNF_Cardinal_Arithmetic.thy
|
blanchet@55098
|
1262 |
Cardinals/Cardinal_Order_Relation_Base.thy ~> BNF_Cardinal_Order_Relation.thy
|
blanchet@55098
|
1263 |
Cardinals/Constructions_on_Wellorders_Base.thy ~> BNF_Constructions_on_Wellorders.thy
|
blanchet@55098
|
1264 |
Cardinals/Wellorder_Embedding_Base.thy ~> BNF_Wellorder_Embedding.thy
|
blanchet@55098
|
1265 |
Cardinals/Wellorder_Relation_Base.thy ~> BNF_Wellorder_Relation.thy
|
blanchet@55098
|
1266 |
BNF/Ctr_Sugar.thy ~> Ctr_Sugar.thy
|
blanchet@55098
|
1267 |
BNF/Basic_BNFs.thy ~> Basic_BNFs.thy
|
blanchet@55098
|
1268 |
BNF/BNF_Comp.thy ~> BNF_Comp.thy
|
blanchet@55098
|
1269 |
BNF/BNF_Def.thy ~> BNF_Def.thy
|
blanchet@55098
|
1270 |
BNF/BNF_FP_Base.thy ~> BNF_FP_Base.thy
|
blanchet@55098
|
1271 |
BNF/BNF_GFP.thy ~> BNF_GFP.thy
|
blanchet@55098
|
1272 |
BNF/BNF_LFP.thy ~> BNF_LFP.thy
|
blanchet@55098
|
1273 |
BNF/BNF_Util.thy ~> BNF_Util.thy
|
blanchet@55098
|
1274 |
BNF/Coinduction.thy ~> Coinduction.thy
|
blanchet@55098
|
1275 |
BNF/More_BNFs.thy ~> Library/More_BNFs.thy
|
blanchet@55098
|
1276 |
BNF/Countable_Type.thy ~> Library/Countable_Set_Type.thy
|
blanchet@55098
|
1277 |
BNF/Examples/* ~> BNF_Examples/*
|
wenzelm@57452
|
1278 |
|
blanchet@55098
|
1279 |
New theories:
|
blanchet@55098
|
1280 |
Wellorder_Extension.thy (split from Zorn.thy)
|
blanchet@55098
|
1281 |
Library/Cardinal_Notations.thy
|
traytel@56942
|
1282 |
Library/BNF_Axomatization.thy
|
blanchet@55098
|
1283 |
BNF_Examples/Misc_Primcorec.thy
|
blanchet@55098
|
1284 |
BNF_Examples/Stream_Processor.thy
|
wenzelm@57452
|
1285 |
|
blanchet@55519
|
1286 |
Discontinued theories:
|
blanchet@55098
|
1287 |
BNF/BNF.thy
|
blanchet@55098
|
1288 |
BNF/Equiv_Relations_More.thy
|
wenzelm@57452
|
1289 |
|
wenzelm@57452
|
1290 |
INCOMPATIBILITY.
|
blanchet@55098
|
1291 |
|
blanchet@56118
|
1292 |
* New (co)datatype package:
|
wenzelm@57452
|
1293 |
- Command 'primcorec' is fully implemented.
|
wenzelm@57452
|
1294 |
- Command 'datatype_new' generates size functions ("size_xxx" and
|
wenzelm@57452
|
1295 |
"size") as required by 'fun'.
|
wenzelm@57452
|
1296 |
- BNFs are integrated with the Lifting tool and new-style
|
wenzelm@57452
|
1297 |
(co)datatypes with Transfer.
|
wenzelm@57452
|
1298 |
- Renamed commands:
|
blanchet@55875
|
1299 |
datatype_new_compat ~> datatype_compat
|
blanchet@55875
|
1300 |
primrec_new ~> primrec
|
blanchet@55875
|
1301 |
wrap_free_constructors ~> free_constructors
|
blanchet@55875
|
1302 |
INCOMPATIBILITY.
|
wenzelm@57452
|
1303 |
- The generated constants "xxx_case" and "xxx_rec" have been renamed
|
blanchet@55875
|
1304 |
"case_xxx" and "rec_xxx" (e.g., "prod_case" ~> "case_prod").
|
blanchet@55875
|
1305 |
INCOMPATIBILITY.
|
wenzelm@57452
|
1306 |
- The constant "xxx_(un)fold" and related theorems are no longer
|
wenzelm@57452
|
1307 |
generated. Use "xxx_(co)rec" or define "xxx_(un)fold" manually
|
wenzelm@57452
|
1308 |
using "prim(co)rec".
|
blanchet@55875
|
1309 |
INCOMPATIBILITY.
|
wenzelm@57452
|
1310 |
- No discriminators are generated for nullary constructors by
|
wenzelm@57452
|
1311 |
default, eliminating the need for the odd "=:" syntax.
|
blanchet@57091
|
1312 |
INCOMPATIBILITY.
|
wenzelm@57452
|
1313 |
- No discriminators or selectors are generated by default by
|
blanchet@57094
|
1314 |
"datatype_new", unless custom names are specified or the new
|
blanchet@57094
|
1315 |
"discs_sels" option is passed.
|
blanchet@57094
|
1316 |
INCOMPATIBILITY.
|
blanchet@55875
|
1317 |
|
blanchet@55643
|
1318 |
* Old datatype package:
|
wenzelm@57452
|
1319 |
- The generated theorems "xxx.cases" and "xxx.recs" have been
|
wenzelm@57452
|
1320 |
renamed "xxx.case" and "xxx.rec" (e.g., "sum.cases" ->
|
wenzelm@57452
|
1321 |
"sum.case"). INCOMPATIBILITY.
|
wenzelm@57452
|
1322 |
- The generated constants "xxx_case", "xxx_rec", and "xxx_size" have
|
wenzelm@57452
|
1323 |
been renamed "case_xxx", "rec_xxx", and "size_xxx" (e.g.,
|
wenzelm@57452
|
1324 |
"prod_case" ~> "case_prod"). INCOMPATIBILITY.
|
wenzelm@57452
|
1325 |
|
wenzelm@57452
|
1326 |
* The types "'a list" and "'a option", their set and map functions,
|
wenzelm@57452
|
1327 |
their relators, and their selectors are now produced using the new
|
wenzelm@57452
|
1328 |
BNF-based datatype package.
|
wenzelm@57452
|
1329 |
|
blanchet@55519
|
1330 |
Renamed constants:
|
blanchet@55519
|
1331 |
Option.set ~> set_option
|
blanchet@55519
|
1332 |
Option.map ~> map_option
|
blanchet@55525
|
1333 |
option_rel ~> rel_option
|
wenzelm@57452
|
1334 |
|
blanchet@55519
|
1335 |
Renamed theorems:
|
blanchet@55585
|
1336 |
set_def ~> set_rec[abs_def]
|
blanchet@55519
|
1337 |
map_def ~> map_rec[abs_def]
|
blanchet@55519
|
1338 |
Option.map_def ~> map_option_case[abs_def] (with "case_option" instead of "rec_option")
|
blanchet@56652
|
1339 |
option.recs ~> option.rec
|
blanchet@55524
|
1340 |
list_all2_def ~> list_all2_iff
|
blanchet@55585
|
1341 |
set.simps ~> set_simps (or the slightly different "list.set")
|
blanchet@55519
|
1342 |
map.simps ~> list.map
|
blanchet@55519
|
1343 |
hd.simps ~> list.sel(1)
|
blanchet@55519
|
1344 |
tl.simps ~> list.sel(2-3)
|
blanchet@55519
|
1345 |
the.simps ~> option.sel
|
wenzelm@57452
|
1346 |
|
wenzelm@57452
|
1347 |
INCOMPATIBILITY.
|
blanchet@55519
|
1348 |
|
blanchet@55933
|
1349 |
* The following map functions and relators have been renamed:
|
blanchet@55939
|
1350 |
sum_map ~> map_sum
|
blanchet@55939
|
1351 |
map_pair ~> map_prod
|
blanchet@55944
|
1352 |
prod_rel ~> rel_prod
|
blanchet@55943
|
1353 |
sum_rel ~> rel_sum
|
blanchet@55945
|
1354 |
fun_rel ~> rel_fun
|
blanchet@55942
|
1355 |
set_rel ~> rel_set
|
blanchet@55942
|
1356 |
filter_rel ~> rel_filter
|
wenzelm@57452
|
1357 |
fset_rel ~> rel_fset (in "src/HOL/Library/FSet.thy")
|
wenzelm@57452
|
1358 |
cset_rel ~> rel_cset (in "src/HOL/Library/Countable_Set_Type.thy")
|
wenzelm@57452
|
1359 |
vset ~> rel_vset (in "src/HOL/Library/Quotient_Set.thy")
|
wenzelm@57452
|
1360 |
|
wenzelm@57452
|
1361 |
INCOMPATIBILITY.
|
wenzelm@57452
|
1362 |
|
kuncar@57826
|
1363 |
* Lifting and Transfer:
|
kuncar@57826
|
1364 |
- a type variable as a raw type is supported
|
kuncar@57826
|
1365 |
- stronger reflexivity prover
|
kuncar@57826
|
1366 |
- rep_eq is always generated by lift_definition
|
wenzelm@57856
|
1367 |
- setup for Lifting/Transfer is now automated for BNFs
|
kuncar@57826
|
1368 |
+ holds for BNFs that do not contain a dead variable
|
wenzelm@57856
|
1369 |
+ relator_eq, relator_mono, relator_distr, relator_domain,
|
kuncar@57826
|
1370 |
relator_eq_onp, quot_map, transfer rules for bi_unique, bi_total,
|
kuncar@57826
|
1371 |
right_unique, right_total, left_unique, left_total are proved
|
kuncar@57826
|
1372 |
automatically
|
kuncar@57826
|
1373 |
+ definition of a predicator is generated automatically
|
kuncar@57826
|
1374 |
+ simplification rules for a predicator definition are proved
|
kuncar@57826
|
1375 |
automatically for datatypes
|
kuncar@57826
|
1376 |
- consolidation of the setup of Lifting/Transfer
|
wenzelm@57856
|
1377 |
+ property that a relator preservers reflexivity is not needed any
|
kuncar@57826
|
1378 |
more
|
kuncar@57826
|
1379 |
Minor INCOMPATIBILITY.
|
wenzelm@57856
|
1380 |
+ left_total and left_unique rules are now transfer rules
|
kuncar@57826
|
1381 |
(reflexivity_rule attribute not needed anymore)
|
kuncar@57826
|
1382 |
INCOMPATIBILITY.
|
wenzelm@57856
|
1383 |
+ Domainp does not have to be a separate assumption in
|
kuncar@57826
|
1384 |
relator_domain theorems (=> more natural statement)
|
kuncar@57826
|
1385 |
INCOMPATIBILITY.
|
kuncar@57826
|
1386 |
- registration of code equations is more robust
|
kuncar@57826
|
1387 |
Potential INCOMPATIBILITY.
|
kuncar@57826
|
1388 |
- respectfulness proof obligation is preprocessed to a more readable
|
kuncar@57826
|
1389 |
form
|
kuncar@57826
|
1390 |
Potential INCOMPATIBILITY.
|
kuncar@57826
|
1391 |
- eq_onp is always unfolded in respectfulness proof obligation
|
kuncar@57826
|
1392 |
Potential INCOMPATIBILITY.
|
wenzelm@57856
|
1393 |
- unregister lifting setup for Code_Numeral.integer and
|
kuncar@57826
|
1394 |
Code_Numeral.natural
|
kuncar@57826
|
1395 |
Potential INCOMPATIBILITY.
|
kuncar@57826
|
1396 |
- Lifting.invariant -> eq_onp
|
kuncar@57826
|
1397 |
INCOMPATIBILITY.
|
wenzelm@57856
|
1398 |
|
wenzelm@57508
|
1399 |
* New internal SAT solver "cdclite" that produces models and proof
|
wenzelm@57508
|
1400 |
traces. This solver replaces the internal SAT solvers "enumerate" and
|
wenzelm@57508
|
1401 |
"dpll". Applications that explicitly used one of these two SAT
|
wenzelm@57508
|
1402 |
solvers should use "cdclite" instead. In addition, "cdclite" is now
|
wenzelm@57508
|
1403 |
the default SAT solver for the "sat" and "satx" proof methods and
|
wenzelm@57508
|
1404 |
corresponding tactics; the old default can be restored using "declare
|
wenzelm@57508
|
1405 |
[[sat_solver = zchaff_with_proofs]]". Minor INCOMPATIBILITY.
|
wenzelm@57508
|
1406 |
|
wenzelm@57508
|
1407 |
* SMT module: A new version of the SMT module, temporarily called
|
wenzelm@57508
|
1408 |
"SMT2", uses SMT-LIB 2 and supports recent versions of Z3 (e.g.,
|
wenzelm@57508
|
1409 |
4.3). The new proof method is called "smt2". CVC3 and CVC4 are also
|
wenzelm@57508
|
1410 |
supported as oracles. Yices is no longer supported, because no version
|
wenzelm@57508
|
1411 |
of the solver can handle both SMT-LIB 2 and quantifiers.
|
wenzelm@57508
|
1412 |
|
wenzelm@57508
|
1413 |
* Activation of Z3 now works via "z3_non_commercial" system option
|
wenzelm@57508
|
1414 |
(without requiring restart), instead of former settings variable
|
wenzelm@57508
|
1415 |
"Z3_NON_COMMERCIAL". The option can be edited in Isabelle/jEdit menu
|
wenzelm@57508
|
1416 |
Plugin Options / Isabelle / General.
|
wenzelm@57508
|
1417 |
|
wenzelm@57508
|
1418 |
* Sledgehammer:
|
wenzelm@57508
|
1419 |
- Z3 can now produce Isar proofs.
|
wenzelm@57508
|
1420 |
- MaSh overhaul:
|
blanchet@57532
|
1421 |
. New SML-based learning algorithms eliminate the dependency on
|
wenzelm@57508
|
1422 |
Python and increase performance and reliability.
|
wenzelm@57508
|
1423 |
. MaSh and MeSh are now used by default together with the
|
wenzelm@57508
|
1424 |
traditional MePo (Meng-Paulson) relevance filter. To disable
|
wenzelm@57508
|
1425 |
MaSh, set the "MaSh" system option in Isabelle/jEdit Plugin
|
wenzelm@57508
|
1426 |
Options / Isabelle / General to "none".
|
wenzelm@57508
|
1427 |
- New option:
|
wenzelm@57508
|
1428 |
smt_proofs
|
wenzelm@57508
|
1429 |
- Renamed options:
|
wenzelm@57508
|
1430 |
isar_compress ~> compress
|
wenzelm@57508
|
1431 |
isar_try0 ~> try0
|
wenzelm@57508
|
1432 |
|
wenzelm@57508
|
1433 |
INCOMPATIBILITY.
|
wenzelm@57508
|
1434 |
|
wenzelm@57508
|
1435 |
* Removed solvers remote_cvc3 and remote_z3. Use cvc3 and z3 instead.
|
wenzelm@57508
|
1436 |
|
wenzelm@57508
|
1437 |
* Nitpick:
|
wenzelm@57508
|
1438 |
- Fixed soundness bug whereby mutually recursive datatypes could
|
wenzelm@57508
|
1439 |
take infinite values.
|
wenzelm@57508
|
1440 |
- Fixed soundness bug with low-level number functions such as
|
wenzelm@57508
|
1441 |
"Abs_Integ" and "Rep_Integ".
|
wenzelm@57508
|
1442 |
- Removed "std" option.
|
wenzelm@57508
|
1443 |
- Renamed "show_datatypes" to "show_types" and "hide_datatypes" to
|
wenzelm@57508
|
1444 |
"hide_types".
|
wenzelm@57508
|
1445 |
|
wenzelm@57508
|
1446 |
* Metis: Removed legacy proof method 'metisFT'. Use 'metis
|
wenzelm@57508
|
1447 |
(full_types)' instead. INCOMPATIBILITY.
|
wenzelm@57508
|
1448 |
|
wenzelm@57508
|
1449 |
* Try0: Added 'algebra' and 'meson' to the set of proof methods.
|
wenzelm@57508
|
1450 |
|
wenzelm@57508
|
1451 |
* Adjustion of INF and SUP operations:
|
wenzelm@57508
|
1452 |
- Elongated constants INFI and SUPR to INFIMUM and SUPREMUM.
|
wenzelm@57508
|
1453 |
- Consolidated theorem names containing INFI and SUPR: have INF and
|
wenzelm@57508
|
1454 |
SUP instead uniformly.
|
wenzelm@57508
|
1455 |
- More aggressive normalization of expressions involving INF and Inf
|
wenzelm@57508
|
1456 |
or SUP and Sup.
|
wenzelm@57508
|
1457 |
- INF_image and SUP_image do not unfold composition.
|
wenzelm@57508
|
1458 |
- Dropped facts INF_comp, SUP_comp.
|
wenzelm@57508
|
1459 |
- Default congruence rules strong_INF_cong and strong_SUP_cong, with
|
wenzelm@57508
|
1460 |
simplifier implication in premises. Generalize and replace former
|
wenzelm@57508
|
1461 |
INT_cong, SUP_cong
|
wenzelm@57508
|
1462 |
|
wenzelm@57508
|
1463 |
INCOMPATIBILITY.
|
wenzelm@57508
|
1464 |
|
wenzelm@57508
|
1465 |
* SUP and INF generalized to conditionally_complete_lattice.
|
wenzelm@57508
|
1466 |
|
wenzelm@57508
|
1467 |
* Swapped orientation of facts image_comp and vimage_comp:
|
wenzelm@57508
|
1468 |
|
wenzelm@57508
|
1469 |
image_compose ~> image_comp [symmetric]
|
wenzelm@57508
|
1470 |
image_comp ~> image_comp [symmetric]
|
wenzelm@57508
|
1471 |
vimage_compose ~> vimage_comp [symmetric]
|
wenzelm@57508
|
1472 |
vimage_comp ~> vimage_comp [symmetric]
|
wenzelm@57508
|
1473 |
|
wenzelm@57508
|
1474 |
INCOMPATIBILITY.
|
wenzelm@57508
|
1475 |
|
wenzelm@57504
|
1476 |
* Theory reorganization: split of Big_Operators.thy into
|
wenzelm@57504
|
1477 |
Groups_Big.thy and Lattices_Big.thy.
|
blanchet@55098
|
1478 |
|
haftmann@57418
|
1479 |
* Consolidated some facts about big group operators:
|
haftmann@57418
|
1480 |
|
haftmann@57418
|
1481 |
setsum_0' ~> setsum.neutral
|
haftmann@57418
|
1482 |
setsum_0 ~> setsum.neutral_const
|
haftmann@57418
|
1483 |
setsum_addf ~> setsum.distrib
|
haftmann@57418
|
1484 |
setsum_cartesian_product ~> setsum.cartesian_product
|
haftmann@57418
|
1485 |
setsum_cases ~> setsum.If_cases
|
haftmann@57418
|
1486 |
setsum_commute ~> setsum.commute
|
haftmann@57418
|
1487 |
setsum_cong ~> setsum.cong
|
haftmann@57418
|
1488 |
setsum_delta ~> setsum.delta
|
haftmann@57418
|
1489 |
setsum_delta' ~> setsum.delta'
|
haftmann@57418
|
1490 |
setsum_diff1' ~> setsum.remove
|
haftmann@57418
|
1491 |
setsum_empty ~> setsum.empty
|
haftmann@57418
|
1492 |
setsum_infinite ~> setsum.infinite
|
haftmann@57418
|
1493 |
setsum_insert ~> setsum.insert
|
haftmann@57418
|
1494 |
setsum_inter_restrict'' ~> setsum.inter_filter
|
haftmann@57418
|
1495 |
setsum_mono_zero_cong_left ~> setsum.mono_neutral_cong_left
|
haftmann@57418
|
1496 |
setsum_mono_zero_cong_right ~> setsum.mono_neutral_cong_right
|
haftmann@57418
|
1497 |
setsum_mono_zero_left ~> setsum.mono_neutral_left
|
haftmann@57418
|
1498 |
setsum_mono_zero_right ~> setsum.mono_neutral_right
|
haftmann@57418
|
1499 |
setsum_reindex ~> setsum.reindex
|
haftmann@57418
|
1500 |
setsum_reindex_cong ~> setsum.reindex_cong
|
haftmann@57418
|
1501 |
setsum_reindex_nonzero ~> setsum.reindex_nontrivial
|
haftmann@57418
|
1502 |
setsum_restrict_set ~> setsum.inter_restrict
|
haftmann@57418
|
1503 |
setsum_Plus ~> setsum.Plus
|
haftmann@57418
|
1504 |
setsum_setsum_restrict ~> setsum.commute_restrict
|
haftmann@57418
|
1505 |
setsum_Sigma ~> setsum.Sigma
|
haftmann@57418
|
1506 |
setsum_subset_diff ~> setsum.subset_diff
|
haftmann@57418
|
1507 |
setsum_Un_disjoint ~> setsum.union_disjoint
|
haftmann@57418
|
1508 |
setsum_UN_disjoint ~> setsum.UNION_disjoint
|
haftmann@57418
|
1509 |
setsum_Un_Int ~> setsum.union_inter
|
haftmann@57418
|
1510 |
setsum_Union_disjoint ~> setsum.Union_disjoint
|
haftmann@57418
|
1511 |
setsum_UNION_zero ~> setsum.Union_comp
|
haftmann@57418
|
1512 |
setsum_Un_zero ~> setsum.union_inter_neutral
|
haftmann@57418
|
1513 |
strong_setprod_cong ~> setprod.strong_cong
|
haftmann@57418
|
1514 |
strong_setsum_cong ~> setsum.strong_cong
|
haftmann@57418
|
1515 |
setprod_1' ~> setprod.neutral
|
haftmann@57418
|
1516 |
setprod_1 ~> setprod.neutral_const
|
haftmann@57418
|
1517 |
setprod_cartesian_product ~> setprod.cartesian_product
|
haftmann@57418
|
1518 |
setprod_cong ~> setprod.cong
|
haftmann@57418
|
1519 |
setprod_delta ~> setprod.delta
|
haftmann@57418
|
1520 |
setprod_delta' ~> setprod.delta'
|
haftmann@57418
|
1521 |
setprod_empty ~> setprod.empty
|
haftmann@57418
|
1522 |
setprod_infinite ~> setprod.infinite
|
haftmann@57418
|
1523 |
setprod_insert ~> setprod.insert
|
haftmann@57418
|
1524 |
setprod_mono_one_cong_left ~> setprod.mono_neutral_cong_left
|
haftmann@57418
|
1525 |
setprod_mono_one_cong_right ~> setprod.mono_neutral_cong_right
|
haftmann@57418
|
1526 |
setprod_mono_one_left ~> setprod.mono_neutral_left
|
haftmann@57418
|
1527 |
setprod_mono_one_right ~> setprod.mono_neutral_right
|
haftmann@57418
|
1528 |
setprod_reindex ~> setprod.reindex
|
haftmann@57418
|
1529 |
setprod_reindex_cong ~> setprod.reindex_cong
|
haftmann@57418
|
1530 |
setprod_reindex_nonzero ~> setprod.reindex_nontrivial
|
haftmann@57418
|
1531 |
setprod_Sigma ~> setprod.Sigma
|
haftmann@57418
|
1532 |
setprod_subset_diff ~> setprod.subset_diff
|
haftmann@57418
|
1533 |
setprod_timesf ~> setprod.distrib
|
haftmann@57418
|
1534 |
setprod_Un2 ~> setprod.union_diff2
|
haftmann@57418
|
1535 |
setprod_Un_disjoint ~> setprod.union_disjoint
|
haftmann@57418
|
1536 |
setprod_UN_disjoint ~> setprod.UNION_disjoint
|
haftmann@57418
|
1537 |
setprod_Un_Int ~> setprod.union_inter
|
haftmann@57418
|
1538 |
setprod_Union_disjoint ~> setprod.Union_disjoint
|
haftmann@57418
|
1539 |
setprod_Un_one ~> setprod.union_inter_neutral
|
haftmann@57418
|
1540 |
|
haftmann@57418
|
1541 |
Dropped setsum_cong2 (simple variant of setsum.cong).
|
haftmann@57418
|
1542 |
Dropped setsum_inter_restrict' (simple variant of setsum.inter_restrict)
|
haftmann@57418
|
1543 |
Dropped setsum_reindex_id, setprod_reindex_id
|
haftmann@57418
|
1544 |
(simple variants of setsum.reindex [symmetric], setprod.reindex [symmetric]).
|
haftmann@57418
|
1545 |
|
wenzelm@57452
|
1546 |
INCOMPATIBILITY.
|
wenzelm@57452
|
1547 |
|
haftmann@54864
|
1548 |
* Abolished slightly odd global lattice interpretation for min/max.
|
haftmann@54864
|
1549 |
|
wenzelm@57452
|
1550 |
Fact consolidations:
|
haftmann@54864
|
1551 |
min_max.inf_assoc ~> min.assoc
|
haftmann@54864
|
1552 |
min_max.inf_commute ~> min.commute
|
haftmann@54864
|
1553 |
min_max.inf_left_commute ~> min.left_commute
|
haftmann@54864
|
1554 |
min_max.inf_idem ~> min.idem
|
haftmann@54864
|
1555 |
min_max.inf_left_idem ~> min.left_idem
|
haftmann@54864
|
1556 |
min_max.inf_right_idem ~> min.right_idem
|
haftmann@54864
|
1557 |
min_max.sup_assoc ~> max.assoc
|
haftmann@54864
|
1558 |
min_max.sup_commute ~> max.commute
|
haftmann@54864
|
1559 |
min_max.sup_left_commute ~> max.left_commute
|
haftmann@54864
|
1560 |
min_max.sup_idem ~> max.idem
|
haftmann@54864
|
1561 |
min_max.sup_left_idem ~> max.left_idem
|
haftmann@54864
|
1562 |
min_max.sup_inf_distrib1 ~> max_min_distrib2
|
haftmann@54864
|
1563 |
min_max.sup_inf_distrib2 ~> max_min_distrib1
|
haftmann@54864
|
1564 |
min_max.inf_sup_distrib1 ~> min_max_distrib2
|
haftmann@54864
|
1565 |
min_max.inf_sup_distrib2 ~> min_max_distrib1
|
haftmann@54864
|
1566 |
min_max.distrib ~> min_max_distribs
|
haftmann@54864
|
1567 |
min_max.inf_absorb1 ~> min.absorb1
|
haftmann@54864
|
1568 |
min_max.inf_absorb2 ~> min.absorb2
|
haftmann@54864
|
1569 |
min_max.sup_absorb1 ~> max.absorb1
|
haftmann@54864
|
1570 |
min_max.sup_absorb2 ~> max.absorb2
|
haftmann@54864
|
1571 |
min_max.le_iff_inf ~> min.absorb_iff1
|
haftmann@54864
|
1572 |
min_max.le_iff_sup ~> max.absorb_iff2
|
haftmann@54864
|
1573 |
min_max.inf_le1 ~> min.cobounded1
|
haftmann@54864
|
1574 |
min_max.inf_le2 ~> min.cobounded2
|
haftmann@54864
|
1575 |
le_maxI1, min_max.sup_ge1 ~> max.cobounded1
|
haftmann@54864
|
1576 |
le_maxI2, min_max.sup_ge2 ~> max.cobounded2
|
haftmann@54864
|
1577 |
min_max.le_infI1 ~> min.coboundedI1
|
haftmann@54864
|
1578 |
min_max.le_infI2 ~> min.coboundedI2
|
haftmann@54864
|
1579 |
min_max.le_supI1 ~> max.coboundedI1
|
haftmann@54864
|
1580 |
min_max.le_supI2 ~> max.coboundedI2
|
haftmann@54864
|
1581 |
min_max.less_infI1 ~> min.strict_coboundedI1
|
haftmann@54864
|
1582 |
min_max.less_infI2 ~> min.strict_coboundedI2
|
haftmann@54864
|
1583 |
min_max.less_supI1 ~> max.strict_coboundedI1
|
haftmann@54864
|
1584 |
min_max.less_supI2 ~> max.strict_coboundedI2
|
haftmann@54864
|
1585 |
min_max.inf_mono ~> min.mono
|
haftmann@54864
|
1586 |
min_max.sup_mono ~> max.mono
|
haftmann@54864
|
1587 |
min_max.le_infI, min_max.inf_greatest ~> min.boundedI
|
haftmann@54864
|
1588 |
min_max.le_supI, min_max.sup_least ~> max.boundedI
|
haftmann@54864
|
1589 |
min_max.le_inf_iff ~> min.bounded_iff
|
haftmann@54864
|
1590 |
min_max.le_sup_iff ~> max.bounded_iff
|
haftmann@54864
|
1591 |
|
haftmann@54864
|
1592 |
For min_max.inf_sup_aci, prefer (one of) min.commute, min.assoc,
|
haftmann@54864
|
1593 |
min.left_commute, min.left_idem, max.commute, max.assoc,
|
haftmann@54864
|
1594 |
max.left_commute, max.left_idem directly.
|
haftmann@54864
|
1595 |
|
wenzelm@57452
|
1596 |
For min_max.inf_sup_ord, prefer (one of) min.cobounded1,
|
wenzelm@57452
|
1597 |
min.cobounded2, max.cobounded1m max.cobounded2 directly.
|
haftmann@54864
|
1598 |
|
haftmann@56807
|
1599 |
For min_ac or max_ac, prefer more general collection ac_simps.
|
haftmann@54864
|
1600 |
|
wenzelm@58604
|
1601 |
INCOMPATIBILITY.
|
haftmann@54864
|
1602 |
|
wenzelm@57452
|
1603 |
* Theorem disambiguation Inf_le_Sup (on finite sets) ~>
|
wenzelm@57452
|
1604 |
Inf_fin_le_Sup_fin. INCOMPATIBILITY.
|
haftmann@54745
|
1605 |
|
haftmann@54295
|
1606 |
* Qualified constant names Wellfounded.acc, Wellfounded.accp.
|
haftmann@54295
|
1607 |
INCOMPATIBILITY.
|
haftmann@54295
|
1608 |
|
haftmann@54228
|
1609 |
* Fact generalization and consolidation:
|
haftmann@54228
|
1610 |
neq_one_mod_two, mod_2_not_eq_zero_eq_one_int ~> not_mod_2_eq_0_eq_1
|
wenzelm@57452
|
1611 |
|
wenzelm@57452
|
1612 |
INCOMPATIBILITY.
|
wenzelm@57452
|
1613 |
|
wenzelm@57452
|
1614 |
* Purely algebraic definition of even. Fact generalization and
|
wenzelm@57452
|
1615 |
consolidation:
|
haftmann@54228
|
1616 |
nat_even_iff_2_dvd, int_even_iff_2_dvd ~> even_iff_2_dvd
|
haftmann@54228
|
1617 |
even_zero_(nat|int) ~> even_zero
|
wenzelm@57452
|
1618 |
|
haftmann@54228
|
1619 |
INCOMPATIBILITY.
|
wenzelm@54055
|
1620 |
|
haftmann@54489
|
1621 |
* Abolished neg_numeral.
|
wenzelm@57452
|
1622 |
- Canonical representation for minus one is "- 1".
|
wenzelm@57452
|
1623 |
- Canonical representation for other negative numbers is "- (numeral _)".
|
wenzelm@57452
|
1624 |
- When devising rule sets for number calculation, consider the
|
haftmann@54587
|
1625 |
following canonical cases: 0, 1, numeral _, - 1, - numeral _.
|
wenzelm@57452
|
1626 |
- HOLogic.dest_number also recognizes numerals in non-canonical forms
|
wenzelm@54893
|
1627 |
like "numeral One", "- numeral One", "- 0" and even "- ... - _".
|
wenzelm@57452
|
1628 |
- Syntax for negative numerals is mere input syntax.
|
wenzelm@57452
|
1629 |
|
haftmann@56964
|
1630 |
INCOMPATIBILITY.
|
haftmann@54489
|
1631 |
|
wenzelm@57517
|
1632 |
* Reduced name variants for rules on associativity and commutativity:
|
wenzelm@57517
|
1633 |
|
wenzelm@57517
|
1634 |
add_assoc ~> add.assoc
|
wenzelm@57517
|
1635 |
add_commute ~> add.commute
|
wenzelm@57517
|
1636 |
add_left_commute ~> add.left_commute
|
wenzelm@57517
|
1637 |
mult_assoc ~> mult.assoc
|
wenzelm@57517
|
1638 |
mult_commute ~> mult.commute
|
wenzelm@57517
|
1639 |
mult_left_commute ~> mult.left_commute
|
wenzelm@57517
|
1640 |
nat_add_assoc ~> add.assoc
|
wenzelm@57517
|
1641 |
nat_add_commute ~> add.commute
|
wenzelm@57517
|
1642 |
nat_add_left_commute ~> add.left_commute
|
wenzelm@57517
|
1643 |
nat_mult_assoc ~> mult.assoc
|
wenzelm@57517
|
1644 |
nat_mult_commute ~> mult.commute
|
wenzelm@57517
|
1645 |
eq_assoc ~> iff_assoc
|
wenzelm@57517
|
1646 |
eq_left_commute ~> iff_left_commute
|
wenzelm@57517
|
1647 |
|
wenzelm@57517
|
1648 |
INCOMPATIBILITY.
|
wenzelm@57517
|
1649 |
|
wenzelm@57650
|
1650 |
* Fact collections add_ac and mult_ac are considered old-fashioned.
|
wenzelm@57637
|
1651 |
Prefer ac_simps instead, or specify rules
|
wenzelm@57637
|
1652 |
(add|mult).(assoc|commute|left_commute) individually.
|
wenzelm@57517
|
1653 |
|
haftmann@54230
|
1654 |
* Elimination of fact duplicates:
|
haftmann@54230
|
1655 |
equals_zero_I ~> minus_unique
|
haftmann@54230
|
1656 |
diff_eq_0_iff_eq ~> right_minus_eq
|
traytel@54588
|
1657 |
nat_infinite ~> infinite_UNIV_nat
|
traytel@54588
|
1658 |
int_infinite ~> infinite_UNIV_int
|
wenzelm@57452
|
1659 |
|
haftmann@54230
|
1660 |
INCOMPATIBILITY.
|
haftmann@54230
|
1661 |
|
haftmann@54230
|
1662 |
* Fact name consolidation:
|
haftmann@54230
|
1663 |
diff_def, diff_minus, ab_diff_minus ~> diff_conv_add_uminus
|
haftmann@54250
|
1664 |
minus_le_self_iff ~> neg_less_eq_nonneg
|
haftmann@54250
|
1665 |
le_minus_self_iff ~> less_eq_neg_nonpos
|
haftmann@54250
|
1666 |
neg_less_nonneg ~> neg_less_pos
|
haftmann@54250
|
1667 |
less_minus_self_iff ~> less_neg_neg [simp]
|
wenzelm@57452
|
1668 |
|
haftmann@54230
|
1669 |
INCOMPATIBILITY.
|
haftmann@54230
|
1670 |
|
haftmann@54230
|
1671 |
* More simplification rules on unary and binary minus:
|
haftmann@54230
|
1672 |
add_diff_cancel, add_diff_cancel_left, add_le_same_cancel1,
|
haftmann@54230
|
1673 |
add_le_same_cancel2, add_less_same_cancel1, add_less_same_cancel2,
|
haftmann@54230
|
1674 |
add_minus_cancel, diff_add_cancel, le_add_same_cancel1,
|
haftmann@54230
|
1675 |
le_add_same_cancel2, less_add_same_cancel1, less_add_same_cancel2,
|
wenzelm@57452
|
1676 |
minus_add_cancel, uminus_add_conv_diff. These correspondingly have
|
wenzelm@57452
|
1677 |
been taken away from fact collections algebra_simps and field_simps.
|
wenzelm@57452
|
1678 |
INCOMPATIBILITY.
|
haftmann@54230
|
1679 |
|
haftmann@54230
|
1680 |
To restore proofs, the following patterns are helpful:
|
haftmann@54230
|
1681 |
|
haftmann@54230
|
1682 |
a) Arbitrary failing proof not involving "diff_def":
|
haftmann@54230
|
1683 |
Consider simplification with algebra_simps or field_simps.
|
haftmann@54230
|
1684 |
|
haftmann@54230
|
1685 |
b) Lifting rules from addition to subtraction:
|
wenzelm@54893
|
1686 |
Try with "using <rule for addition> of [... "- _" ...]" by simp".
|
haftmann@54230
|
1687 |
|
haftmann@54230
|
1688 |
c) Simplification with "diff_def": just drop "diff_def".
|
haftmann@54230
|
1689 |
Consider simplification with algebra_simps or field_simps;
|
haftmann@54230
|
1690 |
or the brute way with
|
haftmann@54230
|
1691 |
"simp add: diff_conv_add_uminus del: add_uminus_conv_diff".
|
haftmann@54230
|
1692 |
|
wenzelm@57452
|
1693 |
* Introduce bdd_above and bdd_below in theory
|
wenzelm@57452
|
1694 |
Conditionally_Complete_Lattices, use them instead of explicitly
|
wenzelm@57452
|
1695 |
stating boundedness of sets.
|
wenzelm@57452
|
1696 |
|
wenzelm@57452
|
1697 |
* ccpo.admissible quantifies only over non-empty chains to allow more
|
wenzelm@57452
|
1698 |
syntax-directed proof rules; the case of the empty chain shows up as
|
wenzelm@57452
|
1699 |
additional case in fixpoint induction proofs. INCOMPATIBILITY.
|
hoelzl@54264
|
1700 |
|
hoelzl@56214
|
1701 |
* Removed and renamed theorems in Series:
|
hoelzl@56214
|
1702 |
summable_le ~> suminf_le
|
hoelzl@56214
|
1703 |
suminf_le ~> suminf_le_const
|
hoelzl@56214
|
1704 |
series_pos_le ~> setsum_le_suminf
|
hoelzl@56214
|
1705 |
series_pos_less ~> setsum_less_suminf
|
hoelzl@56214
|
1706 |
suminf_ge_zero ~> suminf_nonneg
|
hoelzl@56214
|
1707 |
suminf_gt_zero ~> suminf_pos
|
hoelzl@56214
|
1708 |
suminf_gt_zero_iff ~> suminf_pos_iff
|
hoelzl@56214
|
1709 |
summable_sumr_LIMSEQ_suminf ~> summable_LIMSEQ
|
hoelzl@56214
|
1710 |
suminf_0_le ~> suminf_nonneg [rotate]
|
hoelzl@56214
|
1711 |
pos_summable ~> summableI_nonneg_bounded
|
hoelzl@56214
|
1712 |
ratio_test ~> summable_ratio_test
|
hoelzl@56214
|
1713 |
|
hoelzl@56214
|
1714 |
removed series_zero, replaced by sums_finite
|
hoelzl@56214
|
1715 |
|
hoelzl@56214
|
1716 |
removed auxiliary lemmas:
|
wenzelm@57452
|
1717 |
|
hoelzl@56214
|
1718 |
sumr_offset, sumr_offset2, sumr_offset3, sumr_offset4, sumr_group,
|
wenzelm@57452
|
1719 |
half, le_Suc_ex_iff, lemma_realpow_diff_sumr,
|
wenzelm@57452
|
1720 |
real_setsum_nat_ivl_bounded, summable_le2, ratio_test_lemma2,
|
wenzelm@57452
|
1721 |
sumr_minus_one_realpow_zerom, sumr_one_lb_realpow_zero,
|
wenzelm@57452
|
1722 |
summable_convergent_sumr_iff, sumr_diff_mult_const
|
wenzelm@57452
|
1723 |
|
hoelzl@56214
|
1724 |
INCOMPATIBILITY.
|
hoelzl@56214
|
1725 |
|
hoelzl@56214
|
1726 |
* Replace (F)DERIV syntax by has_derivative:
|
hoelzl@56214
|
1727 |
- "(f has_derivative f') (at x within s)" replaces "FDERIV f x : s : f'"
|
hoelzl@56214
|
1728 |
|
hoelzl@56214
|
1729 |
- "(f has_field_derivative f') (at x within s)" replaces "DERIV f x : s : f'"
|
hoelzl@56214
|
1730 |
|
hoelzl@56214
|
1731 |
- "f differentiable at x within s" replaces "_ differentiable _ in _" syntax
|
hoelzl@56214
|
1732 |
|
hoelzl@56214
|
1733 |
- removed constant isDiff
|
hoelzl@56214
|
1734 |
|
wenzelm@57452
|
1735 |
- "DERIV f x : f'" and "FDERIV f x : f'" syntax is only available as
|
wenzelm@57452
|
1736 |
input syntax.
|
wenzelm@57452
|
1737 |
|
wenzelm@57452
|
1738 |
- "DERIV f x : s : f'" and "FDERIV f x : s : f'" syntax removed.
|
hoelzl@56214
|
1739 |
|
hoelzl@56214
|
1740 |
- Renamed FDERIV_... lemmas to has_derivative_...
|
hoelzl@56214
|
1741 |
|
hoelzl@56381
|
1742 |
- renamed deriv (the syntax constant used for "DERIV _ _ :> _") to DERIV
|
hoelzl@56381
|
1743 |
|
hoelzl@56381
|
1744 |
- removed DERIV_intros, has_derivative_eq_intros
|
hoelzl@56381
|
1745 |
|
wenzelm@57452
|
1746 |
- introduced derivative_intros and deriative_eq_intros which
|
wenzelm@57452
|
1747 |
includes now rules for DERIV, has_derivative and
|
wenzelm@57452
|
1748 |
has_vector_derivative.
|
hoelzl@56381
|
1749 |
|
hoelzl@56214
|
1750 |
- Other renamings:
|
hoelzl@56214
|
1751 |
differentiable_def ~> real_differentiable_def
|
hoelzl@56214
|
1752 |
differentiableE ~> real_differentiableE
|
hoelzl@56214
|
1753 |
fderiv_def ~> has_derivative_at
|
hoelzl@56214
|
1754 |
field_fderiv_def ~> field_has_derivative_at
|
hoelzl@56214
|
1755 |
isDiff_der ~> differentiable_def
|
hoelzl@56214
|
1756 |
deriv_fderiv ~> has_field_derivative_def
|
hoelzl@56381
|
1757 |
deriv_def ~> DERIV_def
|
wenzelm@57452
|
1758 |
|
wenzelm@57452
|
1759 |
INCOMPATIBILITY.
|
wenzelm@57452
|
1760 |
|
wenzelm@57452
|
1761 |
* Include more theorems in continuous_intros. Remove the
|
wenzelm@57452
|
1762 |
continuous_on_intros, isCont_intros collections, these facts are now
|
wenzelm@57452
|
1763 |
in continuous_intros.
|
wenzelm@57452
|
1764 |
|
wenzelm@57452
|
1765 |
* Theorems about complex numbers are now stated only using Re and Im,
|
wenzelm@57452
|
1766 |
the Complex constructor is not used anymore. It is possible to use
|
wenzelm@57452
|
1767 |
primcorec to defined the behaviour of a complex-valued function.
|
wenzelm@57452
|
1768 |
|
wenzelm@57452
|
1769 |
Removed theorems about the Complex constructor from the simpset, they
|
wenzelm@57452
|
1770 |
are available as the lemma collection legacy_Complex_simps. This
|
wenzelm@57452
|
1771 |
especially removes
|
wenzelm@57452
|
1772 |
|
hoelzl@56889
|
1773 |
i_complex_of_real: "ii * complex_of_real r = Complex 0 r".
|
hoelzl@56889
|
1774 |
|
wenzelm@57452
|
1775 |
Instead the reverse direction is supported with
|
hoelzl@56889
|
1776 |
Complex_eq: "Complex a b = a + \<i> * b"
|
hoelzl@56889
|
1777 |
|
wenzelm@57452
|
1778 |
Moved csqrt from Fundamental_Algebra_Theorem to Complex.
|
hoelzl@56889
|
1779 |
|
hoelzl@56889
|
1780 |
Renamings:
|
hoelzl@56889
|
1781 |
Re/Im ~> complex.sel
|
hoelzl@56889
|
1782 |
complex_Re/Im_zero ~> zero_complex.sel
|
hoelzl@56889
|
1783 |
complex_Re/Im_add ~> plus_complex.sel
|
hoelzl@56889
|
1784 |
complex_Re/Im_minus ~> uminus_complex.sel
|
hoelzl@56889
|
1785 |
complex_Re/Im_diff ~> minus_complex.sel
|
hoelzl@56889
|
1786 |
complex_Re/Im_one ~> one_complex.sel
|
hoelzl@56889
|
1787 |
complex_Re/Im_mult ~> times_complex.sel
|
hoelzl@56889
|
1788 |
complex_Re/Im_inverse ~> inverse_complex.sel
|
hoelzl@56889
|
1789 |
complex_Re/Im_scaleR ~> scaleR_complex.sel
|
hoelzl@56889
|
1790 |
complex_Re/Im_i ~> ii.sel
|
hoelzl@56889
|
1791 |
complex_Re/Im_cnj ~> cnj.sel
|
hoelzl@56889
|
1792 |
Re/Im_cis ~> cis.sel
|
hoelzl@56889
|
1793 |
|
hoelzl@56889
|
1794 |
complex_divide_def ~> divide_complex_def
|
hoelzl@56889
|
1795 |
complex_norm_def ~> norm_complex_def
|
hoelzl@56889
|
1796 |
cmod_def ~> norm_complex_de
|
hoelzl@56889
|
1797 |
|
hoelzl@56889
|
1798 |
Removed theorems:
|
hoelzl@56889
|
1799 |
complex_zero_def
|
hoelzl@56889
|
1800 |
complex_add_def
|
hoelzl@56889
|
1801 |
complex_minus_def
|
hoelzl@56889
|
1802 |
complex_diff_def
|
hoelzl@56889
|
1803 |
complex_one_def
|
hoelzl@56889
|
1804 |
complex_mult_def
|
hoelzl@56889
|
1805 |
complex_inverse_def
|
hoelzl@56889
|
1806 |
complex_scaleR_def
|
hoelzl@56889
|
1807 |
|
wenzelm@57452
|
1808 |
INCOMPATIBILITY.
|
wenzelm@57452
|
1809 |
|
wenzelm@57504
|
1810 |
* Theory Lubs moved HOL image to HOL-Library. It is replaced by
|
wenzelm@57504
|
1811 |
Conditionally_Complete_Lattices. INCOMPATIBILITY.
|
wenzelm@57504
|
1812 |
|
wenzelm@57504
|
1813 |
* HOL-Library: new theory src/HOL/Library/Tree.thy.
|
wenzelm@57504
|
1814 |
|
wenzelm@57504
|
1815 |
* HOL-Library: removed theory src/HOL/Library/Kleene_Algebra.thy; it
|
wenzelm@57504
|
1816 |
is subsumed by session Kleene_Algebra in AFP.
|
wenzelm@57504
|
1817 |
|
wenzelm@57856
|
1818 |
* HOL-Library / theory RBT: various constants and facts are hidden;
|
wenzelm@57856
|
1819 |
lifting setup is unregistered. INCOMPATIBILITY.
|
kuncar@57826
|
1820 |
|
wenzelm@57504
|
1821 |
* HOL-Cardinals: new theory src/HOL/Cardinals/Ordinal_Arithmetic.thy.
|
wenzelm@57504
|
1822 |
|
wenzelm@57504
|
1823 |
* HOL-Word: bit representations prefer type bool over type bit.
|
wenzelm@57504
|
1824 |
INCOMPATIBILITY.
|
wenzelm@57504
|
1825 |
|
wenzelm@57504
|
1826 |
* HOL-Word:
|
wenzelm@57504
|
1827 |
- Abandoned fact collection "word_arith_alts", which is a duplicate
|
wenzelm@57504
|
1828 |
of "word_arith_wis".
|
wenzelm@57504
|
1829 |
- Dropped first (duplicated) element in fact collections
|
wenzelm@57504
|
1830 |
"sint_word_ariths", "word_arith_alts", "uint_word_ariths",
|
wenzelm@57504
|
1831 |
"uint_word_arith_bintrs".
|
wenzelm@57504
|
1832 |
|
wenzelm@57504
|
1833 |
* HOL-Number_Theory:
|
wenzelm@57504
|
1834 |
- consolidated the proofs of the binomial theorem
|
wenzelm@57504
|
1835 |
- the function fib is again of type nat => nat and not overloaded
|
wenzelm@57504
|
1836 |
- no more references to Old_Number_Theory in the HOL libraries
|
wenzelm@57504
|
1837 |
(except the AFP)
|
wenzelm@57504
|
1838 |
|
wenzelm@57504
|
1839 |
INCOMPATIBILITY.
|
wenzelm@57504
|
1840 |
|
immler@54787
|
1841 |
* HOL-Multivariate_Analysis:
|
wenzelm@57452
|
1842 |
- Type class ordered_real_vector for ordered vector spaces.
|
wenzelm@57452
|
1843 |
- New theory Complex_Basic_Analysis defining complex derivatives,
|
lp15@57253
|
1844 |
holomorphic functions, etc., ported from HOL Light's canal.ml.
|
wenzelm@57452
|
1845 |
- Changed order of ordered_euclidean_space to be compatible with
|
immler@54787
|
1846 |
pointwise ordering on products. Therefore instance of
|
immler@54787
|
1847 |
conditionally_complete_lattice and ordered_real_vector.
|
immler@54787
|
1848 |
INCOMPATIBILITY: use box instead of greaterThanLessThan or
|
wenzelm@57452
|
1849 |
explicit set-comprehensions with eucl_less for other (half-)open
|
immler@54787
|
1850 |
intervals.
|
immler@57476
|
1851 |
- removed dependencies on type class ordered_euclidean_space with
|
immler@57476
|
1852 |
introduction of "cbox" on euclidean_space
|
immler@57476
|
1853 |
- renamed theorems:
|
immler@57476
|
1854 |
interval ~> box
|
immler@57476
|
1855 |
mem_interval ~> mem_box
|
immler@57476
|
1856 |
interval_eq_empty ~> box_eq_empty
|
immler@57476
|
1857 |
interval_ne_empty ~> box_ne_empty
|
immler@57476
|
1858 |
interval_sing(1) ~> cbox_sing
|
immler@57476
|
1859 |
interval_sing(2) ~> box_sing
|
immler@57476
|
1860 |
subset_interval_imp ~> subset_box_imp
|
immler@57476
|
1861 |
subset_interval ~> subset_box
|
immler@57476
|
1862 |
open_interval ~> open_box
|
immler@57476
|
1863 |
closed_interval ~> closed_cbox
|
immler@57476
|
1864 |
interior_closed_interval ~> interior_cbox
|
immler@57476
|
1865 |
bounded_closed_interval ~> bounded_cbox
|
immler@57476
|
1866 |
compact_interval ~> compact_cbox
|
immler@57476
|
1867 |
bounded_subset_closed_interval_symmetric ~> bounded_subset_cbox_symmetric
|
immler@57476
|
1868 |
bounded_subset_closed_interval ~> bounded_subset_cbox
|
immler@57476
|
1869 |
mem_interval_componentwiseI ~> mem_box_componentwiseI
|
immler@57476
|
1870 |
convex_box ~> convex_prod
|
immler@57476
|
1871 |
rel_interior_real_interval ~> rel_interior_real_box
|
immler@57476
|
1872 |
convex_interval ~> convex_box
|
immler@57476
|
1873 |
convex_hull_eq_real_interval ~> convex_hull_eq_real_cbox
|
immler@57476
|
1874 |
frechet_derivative_within_closed_interval ~> frechet_derivative_within_cbox
|
immler@57476
|
1875 |
content_closed_interval' ~> content_cbox'
|
immler@57476
|
1876 |
elementary_subset_interval ~> elementary_subset_box
|
immler@57476
|
1877 |
diameter_closed_interval ~> diameter_cbox
|
immler@57476
|
1878 |
frontier_closed_interval ~> frontier_cbox
|
immler@57476
|
1879 |
frontier_open_interval ~> frontier_box
|
immler@57476
|
1880 |
bounded_subset_open_interval_symmetric ~> bounded_subset_box_symmetric
|
immler@57476
|
1881 |
closure_open_interval ~> closure_box
|
immler@57476
|
1882 |
open_closed_interval_convex ~> open_cbox_convex
|
immler@57476
|
1883 |
open_interval_midpoint ~> box_midpoint
|
immler@57476
|
1884 |
content_image_affinity_interval ~> content_image_affinity_cbox
|
immler@57476
|
1885 |
is_interval_interval ~> is_interval_cbox + is_interval_box + is_interval_closed_interval
|
immler@57476
|
1886 |
bounded_interval ~> bounded_closed_interval + bounded_boxes
|
immler@57476
|
1887 |
|
immler@57476
|
1888 |
- respective theorems for intervals over the reals:
|
immler@57476
|
1889 |
content_closed_interval + content_cbox
|
immler@57476
|
1890 |
has_integral + has_integral_real
|
immler@57476
|
1891 |
fine_division_exists + fine_division_exists_real
|
immler@57476
|
1892 |
has_integral_null + has_integral_null_real
|
immler@57476
|
1893 |
tagged_division_union_interval + tagged_division_union_interval_real
|
immler@57476
|
1894 |
has_integral_const + has_integral_const_real
|
immler@57476
|
1895 |
integral_const + integral_const_real
|
immler@57476
|
1896 |
has_integral_bound + has_integral_bound_real
|
immler@57476
|
1897 |
integrable_continuous + integrable_continuous_real
|
immler@57476
|
1898 |
integrable_subinterval + integrable_subinterval_real
|
immler@57476
|
1899 |
has_integral_reflect_lemma + has_integral_reflect_lemma_real
|
immler@57476
|
1900 |
integrable_reflect + integrable_reflect_real
|
immler@57476
|
1901 |
integral_reflect + integral_reflect_real
|
immler@57476
|
1902 |
image_affinity_interval + image_affinity_cbox
|
immler@57476
|
1903 |
image_smult_interval + image_smult_cbox
|
immler@57476
|
1904 |
integrable_const + integrable_const_ivl
|
immler@57476
|
1905 |
integrable_on_subinterval + integrable_on_subcbox
|
immler@57476
|
1906 |
|
hoelzl@56369
|
1907 |
- renamed theorems:
|
hoelzl@56369
|
1908 |
derivative_linear ~> has_derivative_bounded_linear
|
hoelzl@56369
|
1909 |
derivative_is_linear ~> has_derivative_linear
|
hoelzl@56369
|
1910 |
bounded_linear_imp_linear ~> bounded_linear.linear
|
hoelzl@56369
|
1911 |
|
hoelzl@56993
|
1912 |
* HOL-Probability:
|
hoelzl@57825
|
1913 |
- Renamed positive_integral to nn_integral:
|
hoelzl@57825
|
1914 |
|
hoelzl@57825
|
1915 |
. Renamed all lemmas "*positive_integral*" to *nn_integral*"
|
hoelzl@57825
|
1916 |
positive_integral_positive ~> nn_integral_nonneg
|
hoelzl@57825
|
1917 |
|
hoelzl@57825
|
1918 |
. Renamed abbreviation integral\<^sup>P to integral\<^sup>N.
|
hoelzl@57825
|
1919 |
|
wenzelm@57452
|
1920 |
- replaced the Lebesgue integral on real numbers by the more general
|
wenzelm@57452
|
1921 |
Bochner integral for functions into a real-normed vector space.
|
hoelzl@56993
|
1922 |
|
hoelzl@56993
|
1923 |
integral_zero ~> integral_zero / integrable_zero
|
hoelzl@56993
|
1924 |
integral_minus ~> integral_minus / integrable_minus
|
hoelzl@56993
|
1925 |
integral_add ~> integral_add / integrable_add
|
hoelzl@56993
|
1926 |
integral_diff ~> integral_diff / integrable_diff
|
hoelzl@56993
|
1927 |
integral_setsum ~> integral_setsum / integrable_setsum
|
hoelzl@56993
|
1928 |
integral_multc ~> integral_mult_left / integrable_mult_left
|
hoelzl@56993
|
1929 |
integral_cmult ~> integral_mult_right / integrable_mult_right
|
hoelzl@56993
|
1930 |
integral_triangle_inequality~> integral_norm_bound
|
hoelzl@56993
|
1931 |
integrable_nonneg ~> integrableI_nonneg
|
hoelzl@56993
|
1932 |
integral_positive ~> integral_nonneg_AE
|
hoelzl@56993
|
1933 |
integrable_abs_iff ~> integrable_abs_cancel
|
hoelzl@57825
|
1934 |
positive_integral_lim_INF ~> nn_integral_liminf
|
hoelzl@56993
|
1935 |
lebesgue_real_affine ~> lborel_real_affine
|
hoelzl@56993
|
1936 |
borel_integral_has_integral ~> has_integral_lebesgue_integral
|
hoelzl@56993
|
1937 |
integral_indicator ~>
|
hoelzl@56993
|
1938 |
integral_real_indicator / integrable_real_indicator
|
hoelzl@57825
|
1939 |
positive_integral_fst ~> nn_integral_fst'
|
hoelzl@57825
|
1940 |
positive_integral_fst_measurable ~> nn_integral_fst
|
hoelzl@57825
|
1941 |
positive_integral_snd_measurable ~> nn_integral_snd
|
hoelzl@56993
|
1942 |
|
hoelzl@56993
|
1943 |
integrable_fst_measurable ~>
|
hoelzl@56993
|
1944 |
integral_fst / integrable_fst / AE_integrable_fst
|
hoelzl@56993
|
1945 |
|
hoelzl@56993
|
1946 |
integrable_snd_measurable ~>
|
hoelzl@56993
|
1947 |
integral_snd / integrable_snd / AE_integrable_snd
|
hoelzl@56993
|
1948 |
|
hoelzl@56993
|
1949 |
integral_monotone_convergence ~>
|
hoelzl@56993
|
1950 |
integral_monotone_convergence / integrable_monotone_convergence
|
hoelzl@56993
|
1951 |
|
hoelzl@56993
|
1952 |
integral_monotone_convergence_at_top ~>
|
hoelzl@56993
|
1953 |
integral_monotone_convergence_at_top /
|
hoelzl@56993
|
1954 |
integrable_monotone_convergence_at_top
|
hoelzl@56993
|
1955 |
|
hoelzl@56993
|
1956 |
has_integral_iff_positive_integral_lebesgue ~>
|
hoelzl@56993
|
1957 |
has_integral_iff_has_bochner_integral_lebesgue_nonneg
|
hoelzl@56993
|
1958 |
|
hoelzl@56993
|
1959 |
lebesgue_integral_has_integral ~>
|
hoelzl@56993
|
1960 |
has_integral_integrable_lebesgue_nonneg
|
hoelzl@56993
|
1961 |
|
hoelzl@56993
|
1962 |
positive_integral_lebesgue_has_integral ~>
|
hoelzl@56993
|
1963 |
integral_has_integral_lebesgue_nonneg /
|
hoelzl@56993
|
1964 |
integrable_has_integral_lebesgue_nonneg
|
hoelzl@56993
|
1965 |
|
hoelzl@56993
|
1966 |
lebesgue_integral_real_affine ~>
|
hoelzl@57825
|
1967 |
nn_integral_real_affine
|
hoelzl@56993
|
1968 |
|
hoelzl@56993
|
1969 |
has_integral_iff_positive_integral_lborel ~>
|
hoelzl@56993
|
1970 |
integral_has_integral_nonneg / integrable_has_integral_nonneg
|
hoelzl@56993
|
1971 |
|
hoelzl@56993
|
1972 |
The following theorems where removed:
|
hoelzl@56993
|
1973 |
|
hoelzl@56993
|
1974 |
lebesgue_integral_nonneg
|
hoelzl@56993
|
1975 |
lebesgue_integral_uminus
|
hoelzl@56993
|
1976 |
lebesgue_integral_cmult
|
hoelzl@56993
|
1977 |
lebesgue_integral_multc
|
hoelzl@56993
|
1978 |
lebesgue_integral_cmult_nonneg
|
hoelzl@56993
|
1979 |
integral_cmul_indicator
|
hoelzl@56993
|
1980 |
integral_real
|
wenzelm@54672
|
1981 |
|
wenzelm@57452
|
1982 |
- Formalized properties about exponentially, Erlang, and normal
|
wenzelm@57452
|
1983 |
distributed random variables.
|
wenzelm@57452
|
1984 |
|
wenzelm@57504
|
1985 |
* HOL-Decision_Procs: Separate command 'approximate' for approximative
|
wenzelm@57504
|
1986 |
computation in src/HOL/Decision_Procs/Approximation. Minor
|
wenzelm@57504
|
1987 |
INCOMPATIBILITY.
|
wenzelm@57452
|
1988 |
|
nipkow@57112
|
1989 |
|
wenzelm@55622
|
1990 |
*** Scala ***
|
wenzelm@55622
|
1991 |
|
wenzelm@55622
|
1992 |
* The signature and semantics of Document.Snapshot.cumulate_markup /
|
wenzelm@55622
|
1993 |
select_markup have been clarified. Markup is now traversed in the
|
wenzelm@55622
|
1994 |
order of reports given by the prover: later markup is usually more
|
wenzelm@55622
|
1995 |
specific and may override results accumulated so far. The elements
|
wenzelm@55622
|
1996 |
guard is mandatory and checked precisely. Subtle INCOMPATIBILITY.
|
wenzelm@55622
|
1997 |
|
wenzelm@57452
|
1998 |
* Substantial reworking of internal PIDE protocol communication
|
wenzelm@57452
|
1999 |
channels. INCOMPATIBILITY.
|
wenzelm@57452
|
2000 |
|
wenzelm@55622
|
2001 |
|
wenzelm@54449
|
2002 |
*** ML ***
|
wenzelm@54449
|
2003 |
|
wenzelm@57504
|
2004 |
* Subtle change of semantics of Thm.eq_thm: theory stamps are not
|
wenzelm@57504
|
2005 |
compared (according to Thm.thm_ord), but assumed to be covered by the
|
wenzelm@57504
|
2006 |
current background theory. Thus equivalent data produced in different
|
wenzelm@57504
|
2007 |
branches of the theory graph usually coincides (e.g. relevant for
|
wenzelm@57504
|
2008 |
theory merge). Note that the softer Thm.eq_thm_prop is often more
|
wenzelm@57504
|
2009 |
appropriate than Thm.eq_thm.
|
wenzelm@57504
|
2010 |
|
wenzelm@57504
|
2011 |
* Proper context for basic Simplifier operations: rewrite_rule,
|
wenzelm@57504
|
2012 |
rewrite_goals_rule, rewrite_goals_tac etc. INCOMPATIBILITY, need to
|
wenzelm@57504
|
2013 |
pass runtime Proof.context (and ensure that the simplified entity
|
wenzelm@57504
|
2014 |
actually belongs to it).
|
wenzelm@57504
|
2015 |
|
wenzelm@57504
|
2016 |
* Proper context discipline for read_instantiate and instantiate_tac:
|
wenzelm@57504
|
2017 |
variables that are meant to become schematic need to be given as
|
wenzelm@57504
|
2018 |
fixed, and are generalized by the explicit context of local variables.
|
wenzelm@57504
|
2019 |
This corresponds to Isar attributes "where" and "of" with 'for'
|
wenzelm@57504
|
2020 |
declaration. INCOMPATIBILITY, also due to potential change of indices
|
wenzelm@57504
|
2021 |
of schematic variables.
|
wenzelm@57504
|
2022 |
|
wenzelm@56303
|
2023 |
* Moved ML_Compiler.exn_trace and other operations on exceptions to
|
wenzelm@56303
|
2024 |
structure Runtime. Minor INCOMPATIBILITY.
|
wenzelm@56303
|
2025 |
|
wenzelm@56279
|
2026 |
* Discontinued old Toplevel.debug in favour of system option
|
wenzelm@57452
|
2027 |
"ML_exception_trace", which may be also declared within the context
|
wenzelm@57452
|
2028 |
via "declare [[ML_exception_trace = true]]". Minor INCOMPATIBILITY.
|
wenzelm@56279
|
2029 |
|
wenzelm@56281
|
2030 |
* Renamed configuration option "ML_trace" to "ML_source_trace". Minor
|
wenzelm@56281
|
2031 |
INCOMPATIBILITY.
|
wenzelm@56281
|
2032 |
|
wenzelm@56281
|
2033 |
* Configuration option "ML_print_depth" controls the pretty-printing
|
wenzelm@56281
|
2034 |
depth of the ML compiler within the context. The old print_depth in
|
wenzelm@56285
|
2035 |
ML is still available as default_print_depth, but rarely used. Minor
|
wenzelm@56285
|
2036 |
INCOMPATIBILITY.
|
wenzelm@56279
|
2037 |
|
wenzelm@54449
|
2038 |
* Toplevel function "use" refers to raw ML bootstrap environment,
|
wenzelm@54449
|
2039 |
without Isar context nor antiquotations. Potential INCOMPATIBILITY.
|
wenzelm@54449
|
2040 |
Note that 'ML_file' is the canonical command to load ML files into the
|
wenzelm@54449
|
2041 |
formal context.
|
wenzelm@54449
|
2042 |
|
wenzelm@56205
|
2043 |
* Simplified programming interface to define ML antiquotations, see
|
wenzelm@56205
|
2044 |
structure ML_Antiquotation. Minor INCOMPATIBILITY.
|
wenzelm@56069
|
2045 |
|
wenzelm@56071
|
2046 |
* ML antiquotation @{here} refers to its source position, which is
|
wenzelm@56071
|
2047 |
occasionally useful for experimentation and diagnostic purposes.
|
wenzelm@56071
|
2048 |
|
wenzelm@56135
|
2049 |
* ML antiquotation @{path} produces a Path.T value, similarly to
|
wenzelm@56135
|
2050 |
Path.explode, but with compile-time check against the file-system and
|
wenzelm@56135
|
2051 |
some PIDE markup. Note that unlike theory source, ML does not have a
|
wenzelm@56135
|
2052 |
well-defined master directory, so an absolute symbolic path
|
wenzelm@56135
|
2053 |
specification is usually required, e.g. "~~/src/HOL".
|
wenzelm@56135
|
2054 |
|
wenzelm@56399
|
2055 |
* ML antiquotation @{print} inlines a function to print an arbitrary
|
wenzelm@56399
|
2056 |
ML value, which is occasionally useful for diagnostic or demonstration
|
wenzelm@56399
|
2057 |
purposes.
|
wenzelm@56399
|
2058 |
|
wenzelm@54449
|
2059 |
|
wenzelm@54683
|
2060 |
*** System ***
|
wenzelm@54683
|
2061 |
|
wenzelm@57443
|
2062 |
* Proof General with its traditional helper scripts is now an optional
|
wenzelm@57504
|
2063 |
Isabelle component, e.g. see ProofGeneral-4.2-2 from the Isabelle
|
wenzelm@57504
|
2064 |
component repository http://isabelle.in.tum.de/components/. Note that
|
wenzelm@57504
|
2065 |
the "system" manual provides general explanations about add-on
|
wenzelm@57504
|
2066 |
components, especially those that are not bundled with the release.
|
wenzelm@57443
|
2067 |
|
wenzelm@56439
|
2068 |
* The raw Isabelle process executable has been renamed from
|
wenzelm@56439
|
2069 |
"isabelle-process" to "isabelle_process", which conforms to common
|
wenzelm@56439
|
2070 |
shell naming conventions, and allows to define a shell function within
|
wenzelm@56439
|
2071 |
the Isabelle environment to avoid dynamic path lookup. Rare
|
wenzelm@57504
|
2072 |
incompatibility for old tools that do not use the ISABELLE_PROCESS
|
wenzelm@57504
|
2073 |
settings variable.
|
wenzelm@56439
|
2074 |
|
wenzelm@57439
|
2075 |
* Former "isabelle tty" has been superseded by "isabelle console",
|
wenzelm@57439
|
2076 |
with implicit build like "isabelle jedit", and without the mostly
|
wenzelm@57439
|
2077 |
obsolete Isar TTY loop.
|
wenzelm@57439
|
2078 |
|
wenzelm@57452
|
2079 |
* Simplified "isabelle display" tool. Settings variables DVI_VIEWER
|
wenzelm@57452
|
2080 |
and PDF_VIEWER now refer to the actual programs, not shell
|
wenzelm@57452
|
2081 |
command-lines. Discontinued option -c: invocation may be asynchronous
|
wenzelm@57452
|
2082 |
via desktop environment, without any special precautions. Potential
|
wenzelm@57452
|
2083 |
INCOMPATIBILITY with ambitious private settings.
|
wenzelm@57452
|
2084 |
|
wenzelm@57413
|
2085 |
* Removed obsolete "isabelle unsymbolize". Note that the usual format
|
wenzelm@57413
|
2086 |
for email communication is the Unicode rendering of Isabelle symbols,
|
wenzelm@57413
|
2087 |
as produced by Isabelle/jEdit, for example.
|
wenzelm@57413
|
2088 |
|
wenzelm@57452
|
2089 |
* Removed obsolete tool "wwwfind". Similar functionality may be
|
wenzelm@57452
|
2090 |
integrated into Isabelle/jEdit eventually.
|
wenzelm@57452
|
2091 |
|
wenzelm@57452
|
2092 |
* Improved 'display_drafts' concerning desktop integration and
|
wenzelm@57452
|
2093 |
repeated invocation in PIDE front-end: re-use single file
|
wenzelm@57452
|
2094 |
$ISABELLE_HOME_USER/tmp/drafts.pdf and corresponding views.
|
wenzelm@57084
|
2095 |
|
wenzelm@57452
|
2096 |
* Session ROOT specifications require explicit 'document_files' for
|
wenzelm@57452
|
2097 |
robust dependencies on LaTeX sources. Only these explicitly given
|
wenzelm@57452
|
2098 |
files are copied to the document output directory, before document
|
wenzelm@57452
|
2099 |
processing is started.
|
wenzelm@57452
|
2100 |
|
wenzelm@57504
|
2101 |
* Windows: support for regular TeX installation (e.g. MiKTeX) instead
|
wenzelm@57504
|
2102 |
of TeX Live from Cygwin.
|
wenzelm@57504
|
2103 |
|
wenzelm@57084
|
2104 |
|
wenzelm@57693
|
2105 |
|
wenzelm@54639
|
2106 |
New in Isabelle2013-2 (December 2013)
|
wenzelm@54639
|
2107 |
-------------------------------------
|
wenzelm@54639
|
2108 |
|
wenzelm@54639
|
2109 |
*** Prover IDE -- Isabelle/Scala/jEdit ***
|
wenzelm@54639
|
2110 |
|
wenzelm@54639
|
2111 |
* More robust editing of running commands with internal forks,
|
wenzelm@54639
|
2112 |
e.g. non-terminating 'by' steps.
|
wenzelm@54639
|
2113 |
|
wenzelm@54641
|
2114 |
* More relaxed Sledgehammer panel: avoid repeated application of query
|
wenzelm@54641
|
2115 |
after edits surrounding the command location.
|
wenzelm@54641
|
2116 |
|
wenzelm@54648
|
2117 |
* More status information about commands that are interrupted
|
wenzelm@54648
|
2118 |
accidentally (via physical event or Poly/ML runtime system signal,
|
wenzelm@54648
|
2119 |
e.g. out-of-memory).
|
wenzelm@54648
|
2120 |
|
wenzelm@54653
|
2121 |
|
wenzelm@54653
|
2122 |
*** System ***
|
wenzelm@54653
|
2123 |
|
wenzelm@54653
|
2124 |
* More robust termination of external processes managed by
|
wenzelm@54664
|
2125 |
Isabelle/ML: support cancellation of tasks within the range of
|
wenzelm@54664
|
2126 |
milliseconds, as required for PIDE document editing with automatically
|
wenzelm@54664
|
2127 |
tried tools (e.g. Sledgehammer).
|
wenzelm@54653
|
2128 |
|
wenzelm@54648
|
2129 |
* Reactivated Isabelle/Scala kill command for external processes on
|
wenzelm@54648
|
2130 |
Mac OS X, which was accidentally broken in Isabelle2013-1 due to a
|
wenzelm@54648
|
2131 |
workaround for some Debian/Ubuntu Linux versions from 2013.
|
wenzelm@54648
|
2132 |
|
wenzelm@54639
|
2133 |
|
wenzelm@54639
|
2134 |
|
wenzelm@53971
|
2135 |
New in Isabelle2013-1 (November 2013)
|
wenzelm@53971
|
2136 |
-------------------------------------
|
wenzelm@50994
|
2137 |
|
wenzelm@51293
|
2138 |
*** General ***
|
wenzelm@51293
|
2139 |
|
wenzelm@53971
|
2140 |
* Discontinued obsolete 'uses' within theory header. Note that
|
wenzelm@53971
|
2141 |
commands like 'ML_file' work without separate declaration of file
|
wenzelm@53971
|
2142 |
dependencies. Minor INCOMPATIBILITY.
|
wenzelm@53971
|
2143 |
|
wenzelm@53971
|
2144 |
* Discontinued redundant 'use' command, which was superseded by
|
wenzelm@53971
|
2145 |
'ML_file' in Isabelle2013. Minor INCOMPATIBILITY.
|
wenzelm@53971
|
2146 |
|
wenzelm@53016
|
2147 |
* Simplified subscripts within identifiers, using plain \<^sub>
|
wenzelm@53016
|
2148 |
instead of the second copy \<^isub> and \<^isup>. Superscripts are
|
wenzelm@53016
|
2149 |
only for literal tokens within notation; explicit mixfix annotations
|
wenzelm@53016
|
2150 |
for consts or fixed variables may be used as fall-back for unusual
|
wenzelm@53016
|
2151 |
names. Obsolete \<twosuperior> has been expanded to \<^sup>2 in
|
wenzelm@53016
|
2152 |
Isabelle/HOL. INCOMPATIBILITY, use "isabelle update_sub_sup" to
|
wenzelm@53016
|
2153 |
standardize symbols as a starting point for further manual cleanup.
|
wenzelm@53016
|
2154 |
The ML reference variable "legacy_isub_isup" may be set as temporary
|
wenzelm@53016
|
2155 |
workaround, to make the prover accept a subset of the old identifier
|
wenzelm@53016
|
2156 |
syntax.
|
wenzelm@53016
|
2157 |
|
wenzelm@53021
|
2158 |
* Document antiquotations: term style "isub" has been renamed to
|
wenzelm@53021
|
2159 |
"sub". Minor INCOMPATIBILITY.
|
wenzelm@53021
|
2160 |
|
wenzelm@52487
|
2161 |
* Uniform management of "quick_and_dirty" as system option (see also
|
wenzelm@52487
|
2162 |
"isabelle options"), configuration option within the context (see also
|
wenzelm@52487
|
2163 |
Config.get in Isabelle/ML), and attribute in Isabelle/Isar. Minor
|
wenzelm@52487
|
2164 |
INCOMPATIBILITY, need to use more official Isabelle means to access
|
wenzelm@52487
|
2165 |
quick_and_dirty, instead of historical poking into mutable reference.
|
wenzelm@52059
|
2166 |
|
wenzelm@52060
|
2167 |
* Renamed command 'print_configs' to 'print_options'. Minor
|
wenzelm@52060
|
2168 |
INCOMPATIBILITY.
|
wenzelm@52060
|
2169 |
|
wenzelm@52430
|
2170 |
* Proper diagnostic command 'print_state'. Old 'pr' (with its
|
wenzelm@52430
|
2171 |
implicit change of some global references) is retained for now as
|
wenzelm@52430
|
2172 |
control command, e.g. for ProofGeneral 3.7.x.
|
wenzelm@52430
|
2173 |
|
wenzelm@52549
|
2174 |
* Discontinued 'print_drafts' command with its old-fashioned PS output
|
wenzelm@52549
|
2175 |
and Unix command-line print spooling. Minor INCOMPATIBILITY: use
|
wenzelm@52549
|
2176 |
'display_drafts' instead and print via the regular document viewer.
|
wenzelm@52549
|
2177 |
|
wenzelm@53971
|
2178 |
* Updated and extended "isar-ref" and "implementation" manual,
|
wenzelm@53971
|
2179 |
eliminated old "ref" manual.
|
wenzelm@53971
|
2180 |
|
wenzelm@51293
|
2181 |
|
wenzelm@51533
|
2182 |
*** Prover IDE -- Isabelle/Scala/jEdit ***
|
wenzelm@51533
|
2183 |
|
wenzelm@53971
|
2184 |
* New manual "jedit" for Isabelle/jEdit, see isabelle doc or
|
wenzelm@53852
|
2185 |
Documentation panel.
|
wenzelm@53852
|
2186 |
|
wenzelm@53971
|
2187 |
* Dockable window "Documentation" provides access to Isabelle
|
wenzelm@53971
|
2188 |
documentation.
|
wenzelm@52646
|
2189 |
|
wenzelm@52949
|
2190 |
* Dockable window "Find" provides query operations for formal entities
|
wenzelm@52949
|
2191 |
(GUI front-end to 'find_theorems' command).
|
wenzelm@52949
|
2192 |
|
wenzelm@53050
|
2193 |
* Dockable window "Sledgehammer" manages asynchronous / parallel
|
wenzelm@53050
|
2194 |
sledgehammer runs over existing document sources, independently of
|
wenzelm@53050
|
2195 |
normal editing and checking process.
|
wenzelm@53050
|
2196 |
|
wenzelm@51533
|
2197 |
* Dockable window "Timing" provides an overview of relevant command
|
wenzelm@54332
|
2198 |
timing information, depending on option jedit_timing_threshold. The
|
wenzelm@54332
|
2199 |
same timing information is shown in the extended tooltip of the
|
wenzelm@54332
|
2200 |
command keyword, when hovering the mouse over it while the CONTROL or
|
wenzelm@54332
|
2201 |
COMMAND modifier is pressed.
|
wenzelm@51533
|
2202 |
|
wenzelm@53971
|
2203 |
* Improved dockable window "Theories": Continuous checking of proof
|
wenzelm@53971
|
2204 |
document (visible and required parts) may be controlled explicitly,
|
wenzelm@53971
|
2205 |
using check box or shortcut "C+e ENTER". Individual theory nodes may
|
wenzelm@53971
|
2206 |
be marked explicitly as required and checked in full, using check box
|
wenzelm@53971
|
2207 |
or shortcut "C+e SPACE".
|
wenzelm@53971
|
2208 |
|
wenzelm@54305
|
2209 |
* Improved completion mechanism, which is now managed by the
|
wenzelm@54305
|
2210 |
Isabelle/jEdit plugin instead of SideKick. Refined table of Isabelle
|
wenzelm@54305
|
2211 |
symbol abbreviations (see $ISABELLE_HOME/etc/symbols).
|
wenzelm@54305
|
2212 |
|
wenzelm@54319
|
2213 |
* Standard jEdit keyboard shortcut C+b complete-word is remapped to
|
wenzelm@54319
|
2214 |
isabelle.complete for explicit completion in Isabelle sources.
|
wenzelm@54319
|
2215 |
INCOMPATIBILITY wrt. jEdit defaults, may have to invent new shortcuts
|
wenzelm@54319
|
2216 |
to resolve conflict.
|
wenzelm@54319
|
2217 |
|
wenzelm@54305
|
2218 |
* Improved support of various "minor modes" for Isabelle NEWS,
|
wenzelm@54305
|
2219 |
options, session ROOT etc., with completion and SideKick tree view.
|
wenzelm@54305
|
2220 |
|
wenzelm@53971
|
2221 |
* Strictly monotonic document update, without premature cancellation of
|
wenzelm@53971
|
2222 |
running transactions that are still needed: avoid reset/restart of
|
wenzelm@53971
|
2223 |
such command executions while editing.
|
wenzelm@53271
|
2224 |
|
wenzelm@53971
|
2225 |
* Support for asynchronous print functions, as overlay to existing
|
wenzelm@53971
|
2226 |
document content.
|
wenzelm@53971
|
2227 |
|
wenzelm@53971
|
2228 |
* Support for automatic tools in HOL, which try to prove or disprove
|
wenzelm@53971
|
2229 |
toplevel theorem statements.
|
wenzelm@53971
|
2230 |
|
wenzelm@53971
|
2231 |
* Action isabelle.reset-font-size resets main text area font size
|
wenzelm@54365
|
2232 |
according to Isabelle/Scala plugin option "jedit_font_reset_size" (see
|
wenzelm@54365
|
2233 |
also "Plugin Options / Isabelle / General"). It can be bound to some
|
wenzelm@54365
|
2234 |
keyboard shortcut by the user (e.g. C+0 and/or C+NUMPAD0).
|
wenzelm@53971
|
2235 |
|
wenzelm@53971
|
2236 |
* File specifications in jEdit (e.g. file browser) may refer to
|
wenzelm@54351
|
2237 |
$ISABELLE_HOME and $ISABELLE_HOME_USER on all platforms. Discontinued
|
wenzelm@54351
|
2238 |
obsolete $ISABELLE_HOME_WINDOWS variable.
|
wenzelm@53971
|
2239 |
|
wenzelm@53971
|
2240 |
* Improved support for Linux look-and-feel "GTK+", see also "Utilities
|
wenzelm@53971
|
2241 |
/ Global Options / Appearance".
|
wenzelm@53971
|
2242 |
|
wenzelm@53971
|
2243 |
* Improved support of native Mac OS X functionality via "MacOSX"
|
wenzelm@53971
|
2244 |
plugin, which is now enabled by default.
|
wenzelm@53971
|
2245 |
|
wenzelm@51533
|
2246 |
|
wenzelm@51313
|
2247 |
*** Pure ***
|
wenzelm@51313
|
2248 |
|
ballarin@54049
|
2249 |
* Commands 'interpretation' and 'sublocale' are now target-sensitive.
|
ballarin@54049
|
2250 |
In particular, 'interpretation' allows for non-persistent
|
ballarin@54049
|
2251 |
interpretation within "context ... begin ... end" blocks offering a
|
ballarin@54049
|
2252 |
light-weight alternative to 'sublocale'. See "isar-ref" manual for
|
ballarin@54049
|
2253 |
details.
|
haftmann@51747
|
2254 |
|
ballarin@51565
|
2255 |
* Improved locales diagnostic command 'print_dependencies'.
|
ballarin@51565
|
2256 |
|
wenzelm@51313
|
2257 |
* Discontinued obsolete 'axioms' command, which has been marked as
|
wenzelm@51313
|
2258 |
legacy since Isabelle2009-2. INCOMPATIBILITY, use 'axiomatization'
|
wenzelm@51313
|
2259 |
instead, while observing its uniform scope for polymorphism.
|
wenzelm@51313
|
2260 |
|
wenzelm@51316
|
2261 |
* Discontinued empty name bindings in 'axiomatization'.
|
wenzelm@51316
|
2262 |
INCOMPATIBILITY.
|
wenzelm@51316
|
2263 |
|
wenzelm@53971
|
2264 |
* System option "proofs" has been discontinued. Instead the global
|
wenzelm@53971
|
2265 |
state of Proofterm.proofs is persistently compiled into logic images
|
wenzelm@53971
|
2266 |
as required, notably HOL-Proofs. Users no longer need to change
|
wenzelm@53971
|
2267 |
Proofterm.proofs dynamically. Minor INCOMPATIBILITY.
|
wenzelm@53971
|
2268 |
|
wenzelm@53971
|
2269 |
* Syntax translation functions (print_translation etc.) always depend
|
wenzelm@53971
|
2270 |
on Proof.context. Discontinued former "(advanced)" option -- this is
|
wenzelm@53971
|
2271 |
now the default. Minor INCOMPATIBILITY.
|
wenzelm@53971
|
2272 |
|
wenzelm@53971
|
2273 |
* Former global reference trace_unify_fail is now available as
|
wenzelm@53971
|
2274 |
configuration option "unify_trace_failure" (global context only).
|
wenzelm@53971
|
2275 |
|
wenzelm@52463
|
2276 |
* SELECT_GOAL now retains the syntactic context of the overall goal
|
wenzelm@52463
|
2277 |
state (schematic variables etc.). Potential INCOMPATIBILITY in rare
|
wenzelm@52463
|
2278 |
situations.
|
wenzelm@52463
|
2279 |
|
wenzelm@51313
|
2280 |
|
hoelzl@51002
|
2281 |
*** HOL ***
|
hoelzl@51002
|
2282 |
|
wenzelm@54032
|
2283 |
* Stronger precedence of syntax for big intersection and union on
|
wenzelm@54032
|
2284 |
sets, in accordance with corresponding lattice operations.
|
wenzelm@54032
|
2285 |
INCOMPATIBILITY.
|
wenzelm@54032
|
2286 |
|
wenzelm@54032
|
2287 |
* Notation "{p:A. P}" now allows tuple patterns as well.
|
wenzelm@54032
|
2288 |
|
wenzelm@54032
|
2289 |
* Nested case expressions are now translated in a separate check phase
|
wenzelm@54032
|
2290 |
rather than during parsing. The data for case combinators is separated
|
wenzelm@54032
|
2291 |
from the datatype package. The declaration attribute
|
wenzelm@54032
|
2292 |
"case_translation" can be used to register new case combinators:
|
wenzelm@54032
|
2293 |
|
wenzelm@54032
|
2294 |
declare [[case_translation case_combinator constructor1 ... constructorN]]
|
haftmann@52637
|
2295 |
|
haftmann@52435
|
2296 |
* Code generator:
|
wenzelm@53160
|
2297 |
- 'code_printing' unifies 'code_const' / 'code_type' / 'code_class' /
|
wenzelm@53160
|
2298 |
'code_instance'.
|
wenzelm@53160
|
2299 |
- 'code_identifier' declares name hints for arbitrary identifiers in
|
wenzelm@53160
|
2300 |
generated code, subsuming 'code_modulename'.
|
wenzelm@53983
|
2301 |
|
wenzelm@53983
|
2302 |
See the isar-ref manual for syntax diagrams, and the HOL theories for
|
wenzelm@53983
|
2303 |
examples.
|
haftmann@52435
|
2304 |
|
wenzelm@54032
|
2305 |
* Attibute 'code': 'code' now declares concrete and abstract code
|
wenzelm@54032
|
2306 |
equations uniformly. Use explicit 'code equation' and 'code abstract'
|
wenzelm@54032
|
2307 |
to distinguish both when desired.
|
wenzelm@54032
|
2308 |
|
wenzelm@54032
|
2309 |
* Discontinued theories Code_Integer and Efficient_Nat by a more
|
wenzelm@54032
|
2310 |
fine-grain stack of theories Code_Target_Int, Code_Binary_Nat,
|
wenzelm@54032
|
2311 |
Code_Target_Nat and Code_Target_Numeral. See the tutorial on code
|
wenzelm@54032
|
2312 |
generation for details. INCOMPATIBILITY.
|
wenzelm@54032
|
2313 |
|
wenzelm@54032
|
2314 |
* Numeric types are mapped by default to target language numerals:
|
wenzelm@54032
|
2315 |
natural (replaces former code_numeral) and integer (replaces former
|
wenzelm@54032
|
2316 |
code_int). Conversions are available as integer_of_natural /
|
wenzelm@54032
|
2317 |
natural_of_integer / integer_of_nat / nat_of_integer (in HOL) and
|
wenzelm@54032
|
2318 |
Code_Numeral.integer_of_natural / Code_Numeral.natural_of_integer (in
|
wenzelm@54032
|
2319 |
ML). INCOMPATIBILITY.
|
wenzelm@54032
|
2320 |
|
wenzelm@54032
|
2321 |
* Function package: For mutually recursive functions f and g, separate
|
wenzelm@54032
|
2322 |
cases rules f.cases and g.cases are generated instead of unusable
|
wenzelm@54032
|
2323 |
f_g.cases which exposed internal sum types. Potential INCOMPATIBILITY,
|
wenzelm@54032
|
2324 |
in the case that the unusable rule was used nevertheless.
|
wenzelm@54032
|
2325 |
|
wenzelm@54032
|
2326 |
* Function package: For each function f, new rules f.elims are
|
wenzelm@54032
|
2327 |
generated, which eliminate equalities of the form "f x = t".
|
wenzelm@54032
|
2328 |
|
wenzelm@54032
|
2329 |
* New command 'fun_cases' derives ad-hoc elimination rules for
|
wenzelm@54032
|
2330 |
function equations as simplified instances of f.elims, analogous to
|
wenzelm@54032
|
2331 |
inductive_cases. See ~~/src/HOL/ex/Fundefs.thy for some examples.
|
blanchet@53307
|
2332 |
|
kuncar@54021
|
2333 |
* Lifting:
|
kuncar@54021
|
2334 |
- parametrized correspondence relations are now supported:
|
wenzelm@54378
|
2335 |
+ parametricity theorems for the raw term can be specified in
|
kuncar@54021
|
2336 |
the command lift_definition, which allow us to generate stronger
|
kuncar@54021
|
2337 |
transfer rules
|
kuncar@54021
|
2338 |
+ setup_lifting generates stronger transfer rules if parametric
|
kuncar@54021
|
2339 |
correspondence relation can be generated
|
kuncar@54021
|
2340 |
+ various new properties of the relator must be specified to support
|
kuncar@54021
|
2341 |
parametricity
|
kuncar@54021
|
2342 |
+ parametricity theorem for the Quotient relation can be specified
|
kuncar@54021
|
2343 |
- setup_lifting generates domain rules for the Transfer package
|
kuncar@54021
|
2344 |
- stronger reflexivity prover of respectfulness theorems for type
|
kuncar@54021
|
2345 |
copies
|
kuncar@54021
|
2346 |
- ===> and --> are now local. The symbols can be introduced
|
kuncar@54021
|
2347 |
by interpreting the locale lifting_syntax (typically in an
|
kuncar@54021
|
2348 |
anonymous context)
|
wenzelm@54378
|
2349 |
- Lifting/Transfer relevant parts of Library/Quotient_* are now in
|
kuncar@54021
|
2350 |
Main. Potential INCOMPATIBILITY
|
kuncar@54021
|
2351 |
- new commands for restoring and deleting Lifting/Transfer context:
|
kuncar@54021
|
2352 |
lifting_forget, lifting_update
|
wenzelm@54378
|
2353 |
- the command print_quotmaps was renamed to print_quot_maps.
|
kuncar@54021
|
2354 |
INCOMPATIBILITY
|
kuncar@54021
|
2355 |
|
kuncar@54021
|
2356 |
* Transfer:
|
wenzelm@54378
|
2357 |
- better support for domains in Transfer: replace Domainp T
|
kuncar@54021
|
2358 |
by the actual invariant in a transferred goal
|
kuncar@54021
|
2359 |
- transfer rules can have as assumptions other transfer rules
|
kuncar@54021
|
2360 |
- Experimental support for transferring from the raw level to the
|
kuncar@54021
|
2361 |
abstract level: Transfer.transferred attribute
|
kuncar@54021
|
2362 |
- Attribute version of the transfer method: untransferred attribute
|
kuncar@54021
|
2363 |
|
haftmann@52286
|
2364 |
* Reification and reflection:
|
wenzelm@53160
|
2365 |
- Reification is now directly available in HOL-Main in structure
|
wenzelm@53160
|
2366 |
"Reification".
|
wenzelm@53160
|
2367 |
- Reflection now handles multiple lists with variables also.
|
wenzelm@53160
|
2368 |
- The whole reflection stack has been decomposed into conversions.
|
haftmann@52286
|
2369 |
INCOMPATIBILITY.
|
haftmann@52286
|
2370 |
|
haftmann@51489
|
2371 |
* Revised devices for recursive definitions over finite sets:
|
haftmann@51489
|
2372 |
- Only one fundamental fold combinator on finite set remains:
|
haftmann@51489
|
2373 |
Finite_Set.fold :: ('a => 'b => 'b) => 'b => 'a set => 'b
|
haftmann@51489
|
2374 |
This is now identity on infinite sets.
|
wenzelm@52745
|
2375 |
- Locales ("mini packages") for fundamental definitions with
|
haftmann@51489
|
2376 |
Finite_Set.fold: folding, folding_idem.
|
haftmann@51489
|
2377 |
- Locales comm_monoid_set, semilattice_order_set and
|
haftmann@51489
|
2378 |
semilattice_neutr_order_set for big operators on sets.
|
haftmann@51489
|
2379 |
See theory Big_Operators for canonical examples.
|
haftmann@51489
|
2380 |
Note that foundational constants comm_monoid_set.F and
|
haftmann@51489
|
2381 |
semilattice_set.F correspond to former combinators fold_image
|
haftmann@51489
|
2382 |
and fold1 respectively. These are now gone. You may use
|
haftmann@51490
|
2383 |
those foundational constants as substitutes, but it is
|
wenzelm@53983
|
2384 |
preferable to interpret the above locales accordingly.
|
haftmann@51489
|
2385 |
- Dropped class ab_semigroup_idem_mult (special case of lattice,
|
haftmann@51489
|
2386 |
no longer needed in connection with Finite_Set.fold etc.)
|
haftmann@51489
|
2387 |
- Fact renames:
|
haftmann@51489
|
2388 |
card.union_inter ~> card_Un_Int [symmetric]
|
haftmann@51489
|
2389 |
card.union_disjoint ~> card_Un_disjoint
|
haftmann@51489
|
2390 |
INCOMPATIBILITY.
|
haftmann@51489
|
2391 |
|
haftmann@51487
|
2392 |
* Locale hierarchy for abstract orderings and (semi)lattices.
|
haftmann@51487
|
2393 |
|
wenzelm@53526
|
2394 |
* Complete_Partial_Order.admissible is defined outside the type class
|
wenzelm@53526
|
2395 |
ccpo, but with mandatory prefix ccpo. Admissibility theorems lose the
|
wenzelm@53526
|
2396 |
class predicate assumption or sort constraint when possible.
|
Andreas@53362
|
2397 |
INCOMPATIBILITY.
|
Andreas@53362
|
2398 |
|
wenzelm@53160
|
2399 |
* Introduce type class "conditionally_complete_lattice": Like a
|
wenzelm@53160
|
2400 |
complete lattice but does not assume the existence of the top and
|
wenzelm@53160
|
2401 |
bottom elements. Allows to generalize some lemmas about reals and
|
wenzelm@53160
|
2402 |
extended reals. Removed SupInf and replaced it by the instantiation
|
wenzelm@53160
|
2403 |
of conditionally_complete_lattice for real. Renamed lemmas about
|
wenzelm@53160
|
2404 |
conditionally-complete lattice from Sup_... to cSup_... and from
|
wenzelm@53160
|
2405 |
Inf_... to cInf_... to avoid hidding of similar complete lattice
|
wenzelm@53160
|
2406 |
lemmas.
|
wenzelm@53160
|
2407 |
|
wenzelm@53160
|
2408 |
* Introduce type class linear_continuum as combination of
|
wenzelm@53160
|
2409 |
conditionally-complete lattices and inner dense linorders which have
|
wenzelm@53160
|
2410 |
more than one element. INCOMPATIBILITY.
|
wenzelm@53160
|
2411 |
|
wenzelm@53983
|
2412 |
* Introduced type classes order_top and order_bot. The old classes top
|
wenzelm@53983
|
2413 |
and bot only contain the syntax without assumptions. INCOMPATIBILITY:
|
wenzelm@53983
|
2414 |
Rename bot -> order_bot, top -> order_top
|
lammich@53683
|
2415 |
|
wenzelm@53160
|
2416 |
* Introduce type classes "no_top" and "no_bot" for orderings without
|
wenzelm@53160
|
2417 |
top and bottom elements.
|
hoelzl@51732
|
2418 |
|
hoelzl@51732
|
2419 |
* Split dense_linorder into inner_dense_order and no_top, no_bot.
|
hoelzl@51732
|
2420 |
|
hoelzl@51732
|
2421 |
* Complex_Main: Unify and move various concepts from
|
wenzelm@53160
|
2422 |
HOL-Multivariate_Analysis to HOL-Complex_Main.
|
hoelzl@51732
|
2423 |
|
wenzelm@53983
|
2424 |
- Introduce type class (lin)order_topology and
|
wenzelm@53983
|
2425 |
linear_continuum_topology. Allows to generalize theorems about
|
wenzelm@53983
|
2426 |
limits and order. Instances are reals and extended reals.
|
hoelzl@51732
|
2427 |
|
hoelzl@51732
|
2428 |
- continuous and continuos_on from Multivariate_Analysis:
|
wenzelm@53983
|
2429 |
"continuous" is the continuity of a function at a filter. "isCont"
|
wenzelm@53983
|
2430 |
is now an abbrevitation: "isCont x f == continuous (at _) f".
|
wenzelm@53983
|
2431 |
|
wenzelm@53983
|
2432 |
Generalized continuity lemmas from isCont to continuous on an
|
wenzelm@53983
|
2433 |
arbitrary filter.
|
wenzelm@53983
|
2434 |
|
wenzelm@53983
|
2435 |
- compact from Multivariate_Analysis. Use Bolzano's lemma to prove
|
wenzelm@53983
|
2436 |
compactness of closed intervals on reals. Continuous functions
|
wenzelm@53983
|
2437 |
attain infimum and supremum on compact sets. The inverse of a
|
wenzelm@53983
|
2438 |
continuous function is continuous, when the function is continuous
|
wenzelm@53983
|
2439 |
on a compact set.
|
hoelzl@51732
|
2440 |
|
hoelzl@51732
|
2441 |
- connected from Multivariate_Analysis. Use it to prove the
|
hoelzl@51775
|
2442 |
intermediate value theorem. Show connectedness of intervals on
|
hoelzl@51775
|
2443 |
linear_continuum_topology).
|
hoelzl@51732
|
2444 |
|
hoelzl@51732
|
2445 |
- first_countable_topology from Multivariate_Analysis. Is used to
|
wenzelm@53983
|
2446 |
show equivalence of properties on the neighbourhood filter of x and
|
wenzelm@53983
|
2447 |
on all sequences converging to x.
|
wenzelm@53983
|
2448 |
|
wenzelm@53983
|
2449 |
- FDERIV: Definition of has_derivative moved to Deriv.thy. Moved
|
wenzelm@53983
|
2450 |
theorems from Library/FDERIV.thy to Deriv.thy and base the
|
wenzelm@53983
|
2451 |
definition of DERIV on FDERIV. Add variants of DERIV and FDERIV
|
wenzelm@53983
|
2452 |
which are restricted to sets, i.e. to represent derivatives from
|
wenzelm@53983
|
2453 |
left or right.
|
hoelzl@51732
|
2454 |
|
hoelzl@51732
|
2455 |
- Removed the within-filter. It is replaced by the principal filter:
|
hoelzl@51732
|
2456 |
|
hoelzl@51732
|
2457 |
F within X = inf F (principal X)
|
hoelzl@51732
|
2458 |
|
hoelzl@51732
|
2459 |
- Introduce "at x within U" as a single constant, "at x" is now an
|
hoelzl@51732
|
2460 |
abbreviation for "at x within UNIV"
|
hoelzl@51732
|
2461 |
|
wenzelm@53983
|
2462 |
- Introduce named theorem collections tendsto_intros,
|
wenzelm@53983
|
2463 |
continuous_intros, continuous_on_intros and FDERIV_intros. Theorems
|
wenzelm@53983
|
2464 |
in tendsto_intros (or FDERIV_intros) are also available as
|
wenzelm@53983
|
2465 |
tendsto_eq_intros (or FDERIV_eq_intros) where the right-hand side
|
wenzelm@53983
|
2466 |
is replaced by a congruence rule. This allows to apply them as
|
wenzelm@53983
|
2467 |
intro rules and then proving equivalence by the simplifier.
|
hoelzl@51732
|
2468 |
|
hoelzl@51732
|
2469 |
- Restructured theories in HOL-Complex_Main:
|
hoelzl@51732
|
2470 |
|
hoelzl@51732
|
2471 |
+ Moved RealDef and RComplete into Real
|
hoelzl@51732
|
2472 |
|
hoelzl@51732
|
2473 |
+ Introduced Topological_Spaces and moved theorems about
|
hoelzl@51732
|
2474 |
topological spaces, filters, limits and continuity to it
|
hoelzl@51732
|
2475 |
|
hoelzl@51732
|
2476 |
+ Renamed RealVector to Real_Vector_Spaces
|
hoelzl@51732
|
2477 |
|
wenzelm@53983
|
2478 |
+ Split Lim, SEQ, Series into Topological_Spaces,
|
wenzelm@53983
|
2479 |
Real_Vector_Spaces, and Limits
|
hoelzl@51732
|
2480 |
|
hoelzl@51732
|
2481 |
+ Moved Ln and Log to Transcendental
|
hoelzl@51732
|
2482 |
|
hoelzl@51732
|
2483 |
+ Moved theorems about continuity from Deriv to Topological_Spaces
|
hoelzl@51732
|
2484 |
|
hoelzl@51732
|
2485 |
- Remove various auxiliary lemmas.
|
hoelzl@51732
|
2486 |
|
hoelzl@51732
|
2487 |
INCOMPATIBILITY.
|
hoelzl@51002
|
2488 |
|
blanchet@53738
|
2489 |
* Nitpick:
|
blanchet@55889
|
2490 |
- Added option "spy".
|
blanchet@55889
|
2491 |
- Reduce incidence of "too high arity" errors.
|
blanchet@53738
|
2492 |
|
blanchet@51137
|
2493 |
* Sledgehammer:
|
blanchet@51137
|
2494 |
- Renamed option:
|
blanchet@51137
|
2495 |
isar_shrink ~> isar_compress
|
blanchet@53738
|
2496 |
INCOMPATIBILITY.
|
blanchet@55889
|
2497 |
- Added options "isar_try0", "spy".
|
blanchet@55889
|
2498 |
- Better support for "isar_proofs".
|
blanchet@55889
|
2499 |
- MaSh has been fined-tuned and now runs as a local server.
|
blanchet@51137
|
2500 |
|
wenzelm@54032
|
2501 |
* Improved support for ad hoc overloading of constants (see also
|
wenzelm@54032
|
2502 |
isar-ref manual and ~~/src/HOL/ex/Adhoc_Overloading_Examples.thy).
|
wenzelm@54032
|
2503 |
|
wenzelm@54032
|
2504 |
* Library/Polynomial.thy:
|
wenzelm@54032
|
2505 |
- Use lifting for primitive definitions.
|
wenzelm@54032
|
2506 |
- Explicit conversions from and to lists of coefficients, used for
|
wenzelm@54032
|
2507 |
generated code.
|
wenzelm@54032
|
2508 |
- Replaced recursion operator poly_rec by fold_coeffs.
|
wenzelm@54032
|
2509 |
- Prefer pre-existing gcd operation for gcd.
|
wenzelm@54032
|
2510 |
- Fact renames:
|
wenzelm@54032
|
2511 |
poly_eq_iff ~> poly_eq_poly_eq_iff
|
wenzelm@54032
|
2512 |
poly_ext ~> poly_eqI
|
wenzelm@54032
|
2513 |
expand_poly_eq ~> poly_eq_iff
|
wenzelm@54032
|
2514 |
IMCOMPATIBILITY.
|
wenzelm@54032
|
2515 |
|
wenzelm@54032
|
2516 |
* New Library/Simps_Case_Conv.thy: Provides commands simps_of_case and
|
wenzelm@54032
|
2517 |
case_of_simps to convert function definitions between a list of
|
wenzelm@54032
|
2518 |
equations with patterns on the lhs and a single equation with case
|
wenzelm@54032
|
2519 |
expressions on the rhs. See also Ex/Simps_Case_Conv_Examples.thy.
|
wenzelm@54032
|
2520 |
|
wenzelm@54032
|
2521 |
* New Library/FSet.thy: type of finite sets defined as a subtype of
|
wenzelm@54032
|
2522 |
sets defined by Lifting/Transfer.
|
wenzelm@54032
|
2523 |
|
wenzelm@54032
|
2524 |
* Discontinued theory src/HOL/Library/Eval_Witness. INCOMPATIBILITY.
|
wenzelm@54032
|
2525 |
|
wenzelm@54032
|
2526 |
* Consolidation of library theories on product orders:
|
wenzelm@54032
|
2527 |
|
wenzelm@54032
|
2528 |
Product_Lattice ~> Product_Order -- pointwise order on products
|
wenzelm@54032
|
2529 |
Product_ord ~> Product_Lexorder -- lexicographic order on products
|
wenzelm@54032
|
2530 |
|
wenzelm@54032
|
2531 |
INCOMPATIBILITY.
|
wenzelm@54032
|
2532 |
|
wenzelm@53160
|
2533 |
* Imperative-HOL: The MREC combinator is considered legacy and no
|
wenzelm@53160
|
2534 |
longer included by default. INCOMPATIBILITY, use partial_function
|
wenzelm@53160
|
2535 |
instead, or import theory Legacy_Mrec as a fallback.
|
wenzelm@53160
|
2536 |
|
wenzelm@53983
|
2537 |
* HOL-Algebra: Discontinued theories ~~/src/HOL/Algebra/abstract and
|
wenzelm@53983
|
2538 |
~~/src/HOL/Algebra/poly. Existing theories should be based on
|
wenzelm@53983
|
2539 |
~~/src/HOL/Library/Polynomial instead. The latter provides
|
wenzelm@53983
|
2540 |
integration with HOL's type classes for rings. INCOMPATIBILITY.
|
ballarin@51517
|
2541 |
|
wenzelm@54033
|
2542 |
* HOL-BNF:
|
wenzelm@54032
|
2543 |
- Various improvements to BNF-based (co)datatype package, including
|
wenzelm@54032
|
2544 |
new commands "primrec_new", "primcorec", and
|
wenzelm@54032
|
2545 |
"datatype_new_compat", as well as documentation. See
|
wenzelm@54032
|
2546 |
"datatypes.pdf" for details.
|
wenzelm@54032
|
2547 |
- New "coinduction" method to avoid some boilerplate (compared to
|
wenzelm@54032
|
2548 |
coinduct).
|
wenzelm@54032
|
2549 |
- Renamed keywords:
|
wenzelm@54032
|
2550 |
data ~> datatype_new
|
wenzelm@54032
|
2551 |
codata ~> codatatype
|
wenzelm@54032
|
2552 |
bnf_def ~> bnf
|
wenzelm@54032
|
2553 |
- Renamed many generated theorems, including
|
wenzelm@54032
|
2554 |
discs ~> disc
|
wenzelm@54032
|
2555 |
map_comp' ~> map_comp
|
wenzelm@54032
|
2556 |
map_id' ~> map_id
|
wenzelm@54032
|
2557 |
sels ~> sel
|
wenzelm@54032
|
2558 |
set_map' ~> set_map
|
wenzelm@54032
|
2559 |
sets ~> set
|
wenzelm@54032
|
2560 |
IMCOMPATIBILITY.
|
wenzelm@54032
|
2561 |
|
ballarin@51517
|
2562 |
|
wenzelm@51551
|
2563 |
*** ML ***
|
wenzelm@51551
|
2564 |
|
wenzelm@53971
|
2565 |
* Spec_Check is a Quickcheck tool for Isabelle/ML. The ML function
|
wenzelm@53971
|
2566 |
"check_property" allows to check specifications of the form "ALL x y
|
wenzelm@53971
|
2567 |
z. prop x y z". See also ~~/src/Tools/Spec_Check/ with its
|
wenzelm@53971
|
2568 |
Examples.thy in particular.
|
wenzelm@53971
|
2569 |
|
wenzelm@53709
|
2570 |
* Improved printing of exception trace in Poly/ML 5.5.1, with regular
|
wenzelm@53709
|
2571 |
tracing output in the command transaction context instead of physical
|
wenzelm@53709
|
2572 |
stdout. See also Toplevel.debug, Toplevel.debugging and
|
wenzelm@53709
|
2573 |
ML_Compiler.exn_trace.
|
wenzelm@53709
|
2574 |
|
wenzelm@53971
|
2575 |
* ML type "theory" is now immutable, without any special treatment of
|
wenzelm@53971
|
2576 |
drafts or linear updates (which could lead to "stale theory" errors in
|
wenzelm@53971
|
2577 |
the past). Discontinued obsolete operations like Theory.copy,
|
wenzelm@53971
|
2578 |
Theory.checkpoint, and the auxiliary type theory_ref. Minor
|
wenzelm@53971
|
2579 |
INCOMPATIBILITY.
|
wenzelm@53164
|
2580 |
|
wenzelm@51551
|
2581 |
* More uniform naming of goal functions for skipped proofs:
|
wenzelm@51551
|
2582 |
|
wenzelm@51551
|
2583 |
Skip_Proof.prove ~> Goal.prove_sorry
|
wenzelm@51551
|
2584 |
Skip_Proof.prove_global ~> Goal.prove_sorry_global
|
wenzelm@51551
|
2585 |
|
wenzelm@53971
|
2586 |
Minor INCOMPATIBILITY.
|
wenzelm@51703
|
2587 |
|
wenzelm@51717
|
2588 |
* Simplifier tactics and tools use proper Proof.context instead of
|
wenzelm@51717
|
2589 |
historic type simpset. Old-style declarations like addsimps,
|
wenzelm@51717
|
2590 |
addsimprocs etc. operate directly on Proof.context. Raw type simpset
|
wenzelm@51717
|
2591 |
retains its use as snapshot of the main Simplifier context, using
|
wenzelm@51717
|
2592 |
simpset_of and put_simpset on Proof.context. INCOMPATIBILITY -- port
|
wenzelm@51717
|
2593 |
old tools by making them depend on (ctxt : Proof.context) instead of
|
wenzelm@51717
|
2594 |
(ss : simpset), then turn (simpset_of ctxt) into ctxt.
|
wenzelm@51717
|
2595 |
|
wenzelm@53971
|
2596 |
* Modifiers for classical wrappers (e.g. addWrapper, delWrapper)
|
wenzelm@53971
|
2597 |
operate on Proof.context instead of claset, for uniformity with addIs,
|
wenzelm@53971
|
2598 |
addEs, addDs etc. Note that claset_of and put_claset allow to manage
|
wenzelm@53971
|
2599 |
clasets separately from the context.
|
wenzelm@53971
|
2600 |
|
wenzelm@51717
|
2601 |
* Discontinued obsolete ML antiquotations @{claset} and @{simpset}.
|
wenzelm@51717
|
2602 |
INCOMPATIBILITY, use @{context} instead.
|
wenzelm@51717
|
2603 |
|
wenzelm@53971
|
2604 |
* Antiquotation @{theory_context A} is similar to @{theory A}, but
|
wenzelm@53971
|
2605 |
presents the result as initial Proof.context.
|
wenzelm@53971
|
2606 |
|
wenzelm@51551
|
2607 |
|
wenzelm@51398
|
2608 |
*** System ***
|
wenzelm@51398
|
2609 |
|
wenzelm@52052
|
2610 |
* Discontinued obsolete isabelle usedir, mkdir, make -- superseded by
|
wenzelm@52052
|
2611 |
"isabelle build" in Isabelle2013. INCOMPATIBILITY.
|
wenzelm@51398
|
2612 |
|
wenzelm@52054
|
2613 |
* Discontinued obsolete isabelle-process options -f and -u (former
|
wenzelm@52054
|
2614 |
administrative aliases of option -e). Minor INCOMPATIBILITY.
|
wenzelm@52054
|
2615 |
|
wenzelm@52550
|
2616 |
* Discontinued obsolete isabelle print tool, and PRINT_COMMAND
|
wenzelm@52550
|
2617 |
settings variable.
|
wenzelm@52550
|
2618 |
|
wenzelm@52746
|
2619 |
* Discontinued ISABELLE_DOC_FORMAT settings variable and historic
|
wenzelm@52746
|
2620 |
document formats: dvi.gz, ps, ps.gz -- the default document format is
|
wenzelm@52746
|
2621 |
always pdf.
|
wenzelm@52743
|
2622 |
|
wenzelm@52053
|
2623 |
* Isabelle settings variable ISABELLE_BUILD_JAVA_OPTIONS allows to
|
wenzelm@52053
|
2624 |
specify global resources of the JVM process run by isabelle build.
|
wenzelm@52053
|
2625 |
|
wenzelm@52116
|
2626 |
* Toplevel executable $ISABELLE_HOME/bin/isabelle_scala_script allows
|
wenzelm@52116
|
2627 |
to run Isabelle/Scala source files as standalone programs.
|
wenzelm@52116
|
2628 |
|
wenzelm@52439
|
2629 |
* Improved "isabelle keywords" tool (for old-style ProofGeneral
|
wenzelm@52439
|
2630 |
keyword tables): use Isabelle/Scala operations, which inspect outer
|
wenzelm@52439
|
2631 |
syntax without requiring to build sessions first.
|
wenzelm@52439
|
2632 |
|
wenzelm@53971
|
2633 |
* Sessions may be organized via 'chapter' specifications in the ROOT
|
wenzelm@53971
|
2634 |
file, which determines a two-level hierarchy of browser info. The old
|
wenzelm@53971
|
2635 |
tree-like organization via implicit sub-session relation (with its
|
wenzelm@53971
|
2636 |
tendency towards erratic fluctuation of URLs) has been discontinued.
|
wenzelm@53971
|
2637 |
The default chapter is called "Unsorted". Potential INCOMPATIBILITY
|
wenzelm@53971
|
2638 |
for HTML presentation of theories.
|
wenzelm@53971
|
2639 |
|
wenzelm@51398
|
2640 |
|
wenzelm@51398
|
2641 |
|
wenzelm@50993
|
2642 |
New in Isabelle2013 (February 2013)
|
wenzelm@50993
|
2643 |
-----------------------------------
|
wenzelm@47887
|
2644 |
|
wenzelm@47967
|
2645 |
*** General ***
|
wenzelm@47967
|
2646 |
|
wenzelm@50126
|
2647 |
* Theorem status about oracles and unfinished/failed future proofs is
|
wenzelm@50126
|
2648 |
no longer printed by default, since it is incompatible with
|
wenzelm@50126
|
2649 |
incremental / parallel checking of the persistent document model. ML
|
wenzelm@50126
|
2650 |
function Thm.peek_status may be used to inspect a snapshot of the
|
wenzelm@50126
|
2651 |
ongoing evaluation process. Note that in batch mode --- notably
|
wenzelm@50126
|
2652 |
isabelle build --- the system ensures that future proofs of all
|
wenzelm@50126
|
2653 |
accessible theorems in the theory context are finished (as before).
|
wenzelm@50126
|
2654 |
|
wenzelm@49699
|
2655 |
* Configuration option show_markup controls direct inlining of markup
|
wenzelm@49699
|
2656 |
into the printed representation of formal entities --- notably type
|
wenzelm@49699
|
2657 |
and sort constraints. This enables Prover IDE users to retrieve that
|
wenzelm@49699
|
2658 |
information via tooltips in the output window, for example.
|
wenzelm@49699
|
2659 |
|
wenzelm@48890
|
2660 |
* Command 'ML_file' evaluates ML text from a file directly within the
|
wenzelm@48890
|
2661 |
theory, without any predeclaration via 'uses' in the theory header.
|
wenzelm@48890
|
2662 |
|
wenzelm@49243
|
2663 |
* Old command 'use' command and corresponding keyword 'uses' in the
|
wenzelm@49243
|
2664 |
theory header are legacy features and will be discontinued soon.
|
wenzelm@49243
|
2665 |
Tools that load their additional source files may imitate the
|
wenzelm@49243
|
2666 |
'ML_file' implementation, such that the system can take care of
|
wenzelm@49243
|
2667 |
dependencies properly.
|
wenzelm@49243
|
2668 |
|
wenzelm@47967
|
2669 |
* Discontinued obsolete method fastsimp / tactic fast_simp_tac, which
|
wenzelm@47967
|
2670 |
is called fastforce / fast_force_tac already since Isabelle2011-1.
|
wenzelm@47967
|
2671 |
|
wenzelm@50110
|
2672 |
* Updated and extended "isar-ref" and "implementation" manual, reduced
|
wenzelm@50110
|
2673 |
remaining material in old "ref" manual.
|
wenzelm@48120
|
2674 |
|
wenzelm@51050
|
2675 |
* Improved support for auxiliary contexts that indicate block structure
|
wenzelm@51050
|
2676 |
for specifications. Nesting of "context fixes ... context assumes ..."
|
wenzelm@49841
|
2677 |
and "class ... context ...".
|
wenzelm@49841
|
2678 |
|
wenzelm@50772
|
2679 |
* Attribute "consumes" allows a negative value as well, which is
|
wenzelm@50778
|
2680 |
interpreted relatively to the total number of premises of the rule in
|
wenzelm@50772
|
2681 |
the target context. This form of declaration is stable when exported
|
wenzelm@50772
|
2682 |
from a nested 'context' with additional assumptions. It is the
|
wenzelm@50772
|
2683 |
preferred form for definitional packages, notably cases/rules produced
|
wenzelm@50772
|
2684 |
in HOL/inductive and HOL/function.
|
wenzelm@50772
|
2685 |
|
wenzelm@49869
|
2686 |
* More informative error messages for Isar proof commands involving
|
wenzelm@49869
|
2687 |
lazy enumerations (method applications etc.).
|
wenzelm@49869
|
2688 |
|
wenzelm@50213
|
2689 |
* Refined 'help' command to retrieve outer syntax commands according
|
wenzelm@50213
|
2690 |
to name patterns (with clickable results).
|
wenzelm@50213
|
2691 |
|
wenzelm@47967
|
2692 |
|
wenzelm@49968
|
2693 |
*** Prover IDE -- Isabelle/Scala/jEdit ***
|
wenzelm@49968
|
2694 |
|
wenzelm@49968
|
2695 |
* Parallel terminal proofs ('by') are enabled by default, likewise
|
wenzelm@49968
|
2696 |
proofs that are built into packages like 'datatype', 'function'. This
|
wenzelm@49968
|
2697 |
allows to "run ahead" checking the theory specifications on the
|
wenzelm@49968
|
2698 |
surface, while the prover is still crunching on internal
|
wenzelm@49968
|
2699 |
justifications. Unfinished / cancelled proofs are restarted as
|
wenzelm@49968
|
2700 |
required to complete full proof checking eventually.
|
wenzelm@49968
|
2701 |
|
wenzelm@49968
|
2702 |
* Improved output panel with tooltips, hyperlinks etc. based on the
|
wenzelm@49968
|
2703 |
same Rich_Text_Area as regular Isabelle/jEdit buffers. Activation of
|
wenzelm@49968
|
2704 |
tooltips leads to some window that supports the same recursively,
|
wenzelm@49968
|
2705 |
which can lead to stacks of tooltips as the semantic document content
|
wenzelm@49968
|
2706 |
is explored. ESCAPE closes the whole stack, individual windows may be
|
wenzelm@49968
|
2707 |
closed separately, or detached to become independent jEdit dockables.
|
wenzelm@49968
|
2708 |
|
wenzelm@50717
|
2709 |
* Improved support for commands that produce graph output: the text
|
wenzelm@50717
|
2710 |
message contains a clickable area to open a new instance of the graph
|
wenzelm@50717
|
2711 |
browser on demand.
|
wenzelm@50717
|
2712 |
|
wenzelm@49968
|
2713 |
* More robust incremental parsing of outer syntax (partial comments,
|
wenzelm@49968
|
2714 |
malformed symbols). Changing the balance of open/close quotes and
|
wenzelm@49968
|
2715 |
comment delimiters works more conveniently with unfinished situations
|
wenzelm@49968
|
2716 |
that frequently occur in user interaction.
|
wenzelm@49968
|
2717 |
|
wenzelm@49968
|
2718 |
* More efficient painting and improved reactivity when editing large
|
wenzelm@49968
|
2719 |
files. More scalable management of formal document content.
|
wenzelm@49968
|
2720 |
|
wenzelm@50505
|
2721 |
* Smarter handling of tracing messages: prover process pauses after
|
wenzelm@50505
|
2722 |
certain number of messages per command transaction, with some user
|
wenzelm@50505
|
2723 |
dialog to stop or continue. This avoids swamping the front-end with
|
|