# HG changeset patch
# User wenzelm
# Date 1279903366 -7200
# Node ID 94e9302a7fd014e69bbef770fa6a9c5ebf04f890
# Parent 844977c7abebe212e974d9d8f7aeb92ed33ab5f6# Parent 4b7afae88c576c270fc85317a30bf163c5dec1d7
merged
diff -r 844977c7abeb -r 94e9302a7fd0 NEWS
--- a/NEWS Fri Jul 23 10:58:13 2010 +0200
+++ b/NEWS Fri Jul 23 18:42:46 2010 +0200
@@ -14,6 +14,11 @@
consistent view on symbols, while raw explode (or String.explode)
merely give a byte-oriented representation.
+* Special treatment of ML file names has been discontinued.
+Historically, optional extensions .ML or .sml were added on demand --
+at the cost of clarity of file dependencies. Recall that Isabelle/ML
+files exclusively use the .ML extension. Minor INCOMPATIBILTY.
+
*** HOL ***
diff -r 844977c7abeb -r 94e9302a7fd0 lib/html/isabelle.css
--- a/lib/html/isabelle.css Fri Jul 23 10:58:13 2010 +0200
+++ b/lib/html/isabelle.css Fri Jul 23 18:42:46 2010 +0200
@@ -5,8 +5,8 @@
.head { background-color: #FFFFFF; }
.source { background-color: #F0F0F0; padding: 10px; }
-.mlsource { background-color: #F0F0F0; padding: 10px; }
-.mlfooter { background-color: #FFFFFF; }
+.external_source { background-color: #F0F0F0; padding: 10px; }
+.external_footer { background-color: #FFFFFF; }
.theories { background-color: #F0F0F0; padding: 10px; }
.sessions { background-color: #F0F0F0; padding: 10px; }
@@ -14,9 +14,6 @@
.name { font-style: italic; }
.filename { font-family: fixed; }
-/* hide hr for this style */
-hr { height: 0px; border: 0px; }
-
/* basic syntax markup */
diff -r 844977c7abeb -r 94e9302a7fd0 src/HOL/Boogie/Tools/boogie_commands.ML
--- a/src/HOL/Boogie/Tools/boogie_commands.ML Fri Jul 23 10:58:13 2010 +0200
+++ b/src/HOL/Boogie/Tools/boogie_commands.ML Fri Jul 23 18:42:46 2010 +0200
@@ -268,7 +268,7 @@
val _ =
Outer_Syntax.command "boogie_open"
- "Open a new Boogie environment and load a Boogie-generated .b2i file."
+ "open a new Boogie environment and load a Boogie-generated .b2i file"
Keyword.thy_decl
(scan_opt "quiet" -- Parse.name -- vc_offsets >>
(Toplevel.theory o boogie_open))
@@ -296,7 +296,7 @@
val _ =
Outer_Syntax.command "boogie_vc"
- "Enter into proof mode for a specific verification condition."
+ "enter into proof mode for a specific verification condition"
Keyword.thy_goal
(vc_name -- vc_opts >> (fn args =>
(Toplevel.print o Toplevel.theory_to_proof (boogie_vc args))))
@@ -329,7 +329,7 @@
val _ =
Outer_Syntax.improper_command "boogie_status"
- "Show the name and state of all loaded verification conditions."
+ "show the name and state of all loaded verification conditions"
Keyword.diag
(status_test >> status_cmd ||
status_vc >> status_cmd ||
@@ -338,7 +338,7 @@
val _ =
Outer_Syntax.command "boogie_end"
- "Close the current Boogie environment."
+ "close the current Boogie environment"
Keyword.thy_decl
(Scan.succeed (Toplevel.theory boogie_end))
diff -r 844977c7abeb -r 94e9302a7fd0 src/HOL/Tools/Quotient/quotient_info.ML
--- a/src/HOL/Tools/Quotient/quotient_info.ML Fri Jul 23 10:58:13 2010 +0200
+++ b/src/HOL/Tools/Quotient/quotient_info.ML Fri Jul 23 18:42:46 2010 +0200
@@ -282,9 +282,9 @@
Keyword.diag (Scan.succeed (Toplevel.keep (pp_fn o Toplevel.context_of)))
val _ = map improper_command
- [(print_mapsinfo, "print_quotmaps", "prints out all map functions"),
- (print_quotinfo, "print_quotients", "prints out all quotients"),
- (print_qconstinfo, "print_quotconsts", "prints out all quotient constants")]
+ [(print_mapsinfo, "print_quotmaps", "print out all map functions"),
+ (print_quotinfo, "print_quotients", "print out all quotients"),
+ (print_qconstinfo, "print_quotconsts", "print out all quotient constants")]
end; (* structure *)
diff -r 844977c7abeb -r 94e9302a7fd0 src/Pure/Isar/toplevel.ML
--- a/src/Pure/Isar/toplevel.ML Fri Jul 23 10:58:13 2010 +0200
+++ b/src/Pure/Isar/toplevel.ML Fri Jul 23 18:42:46 2010 +0200
@@ -632,7 +632,7 @@
fun run_command thy_name tr st =
(case
(case init_of tr of
- SOME name => Exn.capture (fn () => Thy_Load.check_name thy_name name) ()
+ SOME name => Exn.capture (fn () => Thy_Load.consistent_name thy_name name) ()
| NONE => Exn.Result ()) of
Exn.Result () =>
let val int = is_some (init_of tr) in
diff -r 844977c7abeb -r 94e9302a7fd0 src/Pure/ProofGeneral/pgip_isabelle.ML
--- a/src/Pure/ProofGeneral/pgip_isabelle.ML Fri Jul 23 10:58:13 2010 +0200
+++ b/src/Pure/ProofGeneral/pgip_isabelle.ML Fri Jul 23 18:42:46 2010 +0200
@@ -80,7 +80,7 @@
NONE => (NONE, NONE)
| SOME fname =>
let val path = Path.explode fname in
- case Thy_Load.check_file Path.current path of
+ case Thy_Load.test_file Path.current path of
SOME _ => (SOME (PgipTypes.pgipurl_of_path path), NONE)
| NONE => (NONE, SOME fname)
end);
diff -r 844977c7abeb -r 94e9302a7fd0 src/Pure/Thy/html.ML
--- a/src/Pure/Thy/html.ML Fri Jul 23 10:58:13 2010 +0200
+++ b/src/Pure/Thy/html.ML Fri Jul 23 18:42:46 2010 +0200
@@ -30,7 +30,7 @@
val theory_source: text -> text
val begin_theory: Url.T * string -> string -> (Url.T option * string) list ->
(Url.T * Url.T * bool) list -> text -> text
- val ml_file: Url.T -> string -> text
+ val external_file: Url.T -> string -> text
end;
structure HTML: HTML =
@@ -309,7 +309,7 @@
begin_document ("Index of " ^ session) ^ up_link up ^
para ("View " ^ href_path graph "theory dependencies" ^
implode (map (fn (p, name) => "
\nView " ^ href_path p name) docs)) ^
- "\n\n
" - "\n
" + "\n"; local @@ -379,13 +379,13 @@ end; -(* ML file *) +(* external file *) -fun ml_file path str = +fun external_file path str = begin_document ("File " ^ Url.implode path) ^ - "\n