# HG changeset patch # User wenzelm # Date 1261234946 -3600 # Node ID 3dcb46ae61857993dcad7c8f29292d51b2b1ed81 # Parent 63dd95e3b39325e2cbc67375726326d2308de66f added basic library -- Scala version; added extra support for exceptions -- Scala version; moved exn.ML to accompany exn.scala; diff -r 63dd95e3b393 -r 3dcb46ae6185 src/Pure/General/exn.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/Pure/General/exn.ML Sat Dec 19 16:02:26 2009 +0100 @@ -0,0 +1,61 @@ +(* Title: Pure/General/exn.ML + Author: Makarius + +Extra support for exceptions. +*) + +signature EXN = +sig + datatype 'a result = Result of 'a | Exn of exn + val get_result: 'a result -> 'a option + val get_exn: 'a result -> exn option + val capture: ('a -> 'b) -> 'a -> 'b result + val release: 'a result -> 'a + exception Interrupt + exception EXCEPTIONS of exn list + val flatten: exn -> exn list + val flatten_list: exn list -> exn list + val release_all: 'a result list -> 'a list + val release_first: 'a result list -> 'a list +end; + +structure Exn: EXN = +struct + +(* runtime exceptions as values *) + +datatype 'a result = + Result of 'a | + Exn of exn; + +fun get_result (Result x) = SOME x + | get_result _ = NONE; + +fun get_exn (Exn exn) = SOME exn + | get_exn _ = NONE; + +fun capture f x = Result (f x) handle e => Exn e; + +fun release (Result y) = y + | release (Exn e) = reraise e; + + +(* interrupt and nested exceptions *) + +exception Interrupt = Interrupt; +exception EXCEPTIONS of exn list; + +fun flatten Interrupt = [] + | flatten (EXCEPTIONS exns) = flatten_list exns + | flatten exn = [exn] +and flatten_list exns = List.concat (map flatten exns); + +fun release_all results = + if List.all (fn Result _ => true | _ => false) results + then map (fn Result x => x) results + else raise EXCEPTIONS (flatten_list (List.mapPartial get_exn results)); + +fun release_first results = release_all results + handle EXCEPTIONS (exn :: _) => reraise exn; + +end; diff -r 63dd95e3b393 -r 3dcb46ae6185 src/Pure/General/exn.scala --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/Pure/General/exn.scala Sat Dec 19 16:02:26 2009 +0100 @@ -0,0 +1,28 @@ +/* Title: Pure/General/exn.scala + Author: Makarius + +Extra support for exceptions. +*/ + +package isabelle + + +object Exn +{ + /* runtime exceptions as values */ + + sealed abstract class Result[A] + case class Res[A](val result: A) extends Result[A] + case class Exn[A](val exn: Exception) extends Result[A] + + def capture[A](e: => A): Result[A] = + try { Res(e) } + catch { case exn: RuntimeException => Exn[A](exn) } + + def release[A](result: Result[A]): A = + result match { + case Res(x) => x + case Exn(exn) => throw exn + } +} + diff -r 63dd95e3b393 -r 3dcb46ae6185 src/Pure/IsaMakefile --- a/src/Pure/IsaMakefile Sat Dec 19 11:48:11 2009 +0100 +++ b/src/Pure/IsaMakefile Sat Dec 19 16:02:26 2009 +0100 @@ -23,7 +23,7 @@ BOOTSTRAP_FILES = ML-Systems/compiler_polyml-5.0.ML \ ML-Systems/compiler_polyml-5.2.ML ML-Systems/compiler_polyml-5.3.ML \ - ML-Systems/exn.ML ML-Systems/ml_name_space.ML \ + ML-Systems/ml_name_space.ML \ ML-Systems/ml_pretty.ML ML-Systems/mosml.ML \ ML-Systems/multithreading.ML ML-Systems/multithreading_polyml.ML \ ML-Systems/overloading_smlnj.ML ML-Systems/polyml-5.0.ML \ @@ -51,38 +51,38 @@ Concurrent/synchronized_sequential.ML Concurrent/task_queue.ML \ General/alist.ML General/antiquote.ML General/balanced_tree.ML \ General/basics.ML General/binding.ML General/buffer.ML \ - General/file.ML General/graph.ML General/heap.ML General/integer.ML \ - General/long_name.ML General/markup.ML General/name_space.ML \ - General/ord_list.ML General/output.ML General/path.ML \ - General/position.ML General/pretty.ML General/print_mode.ML \ - General/properties.ML General/queue.ML General/same.ML \ - General/scan.ML General/secure.ML General/seq.ML General/source.ML \ - General/stack.ML General/symbol.ML General/symbol_pos.ML \ - General/table.ML General/url.ML General/xml.ML General/yxml.ML \ - Isar/args.ML Isar/attrib.ML Isar/auto_bind.ML Isar/calculation.ML \ - Isar/class.ML Isar/class_target.ML Isar/code.ML Isar/constdefs.ML \ - Isar/context_rules.ML Isar/element.ML Isar/expression.ML \ - Isar/isar_cmd.ML Isar/isar_document.ML Isar/isar_syn.ML \ - Isar/local_defs.ML Isar/local_syntax.ML Isar/local_theory.ML \ - Isar/locale.ML Isar/method.ML Isar/object_logic.ML Isar/obtain.ML \ - Isar/outer_keyword.ML Isar/outer_lex.ML Isar/outer_parse.ML \ - Isar/outer_syntax.ML Isar/overloading.ML Isar/proof.ML \ - Isar/proof_context.ML Isar/proof_display.ML Isar/proof_node.ML \ - Isar/rule_cases.ML Isar/rule_insts.ML Isar/runtime.ML \ - Isar/skip_proof.ML Isar/spec_parse.ML Isar/spec_rules.ML \ - Isar/specification.ML Isar/theory_target.ML Isar/toplevel.ML \ - Isar/value_parse.ML ML/ml_antiquote.ML ML/ml_compiler.ML \ - ML/ml_compiler_polyml-5.3.ML ML/ml_context.ML ML/ml_env.ML \ - ML/ml_lex.ML ML/ml_parse.ML ML/ml_syntax.ML ML/ml_thms.ML \ - ML-Systems/install_pp_polyml.ML ML-Systems/install_pp_polyml-5.3.ML \ - ML-Systems/use_context.ML Proof/extraction.ML \ - Proof/proof_rewrite_rules.ML Proof/proof_syntax.ML \ - Proof/proofchecker.ML Proof/reconstruct.ML ProofGeneral/pgip.ML \ - ProofGeneral/pgip_input.ML ProofGeneral/pgip_isabelle.ML \ - ProofGeneral/pgip_markup.ML ProofGeneral/pgip_output.ML \ - ProofGeneral/pgip_parser.ML ProofGeneral/pgip_tests.ML \ - ProofGeneral/pgip_types.ML ProofGeneral/preferences.ML \ - ProofGeneral/proof_general_emacs.ML \ + General/exn.ML General/file.ML General/graph.ML General/heap.ML \ + General/integer.ML General/long_name.ML General/markup.ML \ + General/name_space.ML General/ord_list.ML General/output.ML \ + General/path.ML General/position.ML General/pretty.ML \ + General/print_mode.ML General/properties.ML General/queue.ML \ + General/same.ML General/scan.ML General/secure.ML General/seq.ML \ + General/source.ML General/stack.ML General/symbol.ML \ + General/symbol_pos.ML General/table.ML General/url.ML General/xml.ML \ + General/yxml.ML Isar/args.ML Isar/attrib.ML Isar/auto_bind.ML \ + Isar/calculation.ML Isar/class.ML Isar/class_target.ML Isar/code.ML \ + Isar/constdefs.ML Isar/context_rules.ML Isar/element.ML \ + Isar/expression.ML Isar/isar_cmd.ML Isar/isar_document.ML \ + Isar/isar_syn.ML Isar/local_defs.ML Isar/local_syntax.ML \ + Isar/local_theory.ML Isar/locale.ML Isar/method.ML \ + Isar/object_logic.ML Isar/obtain.ML Isar/outer_keyword.ML \ + Isar/outer_lex.ML Isar/outer_parse.ML Isar/outer_syntax.ML \ + Isar/overloading.ML Isar/proof.ML Isar/proof_context.ML \ + Isar/proof_display.ML Isar/proof_node.ML Isar/rule_cases.ML \ + Isar/rule_insts.ML Isar/runtime.ML Isar/skip_proof.ML \ + Isar/spec_parse.ML Isar/spec_rules.ML Isar/specification.ML \ + Isar/theory_target.ML Isar/toplevel.ML Isar/value_parse.ML \ + ML/ml_antiquote.ML ML/ml_compiler.ML ML/ml_compiler_polyml-5.3.ML \ + ML/ml_context.ML ML/ml_env.ML ML/ml_lex.ML ML/ml_parse.ML \ + ML/ml_syntax.ML ML/ml_thms.ML ML-Systems/install_pp_polyml.ML \ + ML-Systems/install_pp_polyml-5.3.ML ML-Systems/use_context.ML \ + Proof/extraction.ML Proof/proof_rewrite_rules.ML \ + Proof/proof_syntax.ML Proof/proofchecker.ML Proof/reconstruct.ML \ + ProofGeneral/pgip.ML ProofGeneral/pgip_input.ML \ + ProofGeneral/pgip_isabelle.ML ProofGeneral/pgip_markup.ML \ + ProofGeneral/pgip_output.ML ProofGeneral/pgip_parser.ML \ + ProofGeneral/pgip_tests.ML ProofGeneral/pgip_types.ML \ + ProofGeneral/preferences.ML ProofGeneral/proof_general_emacs.ML \ ProofGeneral/proof_general_pgip.ML Pure.thy ROOT.ML Syntax/ast.ML \ Syntax/lexicon.ML Syntax/mixfix.ML Syntax/parser.ML \ Syntax/printer.ML Syntax/simple_syntax.ML Syntax/syn_ext.ML \ @@ -121,14 +121,15 @@ ## Scala material -SCALA_FILES = General/event_bus.scala General/linear_set.scala \ - General/markup.scala General/position.scala General/scan.scala \ - General/swing_thread.scala General/symbol.scala General/xml.scala \ - General/yxml.scala Isar/isar_document.scala Isar/outer_keyword.scala \ - System/cygwin.scala System/gui_setup.scala \ +SCALA_FILES = General/event_bus.scala General/exn.scala \ + General/linear_set.scala General/markup.scala General/position.scala \ + General/scan.scala General/swing_thread.scala General/symbol.scala \ + General/xml.scala General/yxml.scala Isar/isar_document.scala \ + Isar/outer_keyword.scala System/cygwin.scala System/gui_setup.scala \ System/isabelle_process.scala System/isabelle_syntax.scala \ System/isabelle_system.scala System/platform.scala \ - Thy/completion.scala Thy/html.scala Thy/thy_header.scala + Thy/completion.scala Thy/html.scala Thy/thy_header.scala \ + library.scala JAR_DIR = $(ISABELLE_HOME)/lib/classes PURE_JAR = $(JAR_DIR)/Pure.jar diff -r 63dd95e3b393 -r 3dcb46ae6185 src/Pure/ML-Systems/exn.ML --- a/src/Pure/ML-Systems/exn.ML Sat Dec 19 11:48:11 2009 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,61 +0,0 @@ -(* Title: Pure/ML-Systems/exn.ML - Author: Makarius - -Extra support for exceptions. -*) - -signature EXN = -sig - datatype 'a result = Exn of exn | Result of 'a - val get_result: 'a result -> 'a option - val get_exn: 'a result -> exn option - val capture: ('a -> 'b) -> 'a -> 'b result - val release: 'a result -> 'a - exception Interrupt - exception EXCEPTIONS of exn list - val flatten: exn -> exn list - val flatten_list: exn list -> exn list - val release_all: 'a result list -> 'a list - val release_first: 'a result list -> 'a list -end; - -structure Exn: EXN = -struct - -(* runtime exceptions as values *) - -datatype 'a result = - Result of 'a | - Exn of exn; - -fun get_result (Result x) = SOME x - | get_result _ = NONE; - -fun get_exn (Exn exn) = SOME exn - | get_exn _ = NONE; - -fun capture f x = Result (f x) handle e => Exn e; - -fun release (Result y) = y - | release (Exn e) = reraise e; - - -(* interrupt and nested exceptions *) - -exception Interrupt = Interrupt; -exception EXCEPTIONS of exn list; - -fun flatten Interrupt = [] - | flatten (EXCEPTIONS exns) = flatten_list exns - | flatten exn = [exn] -and flatten_list exns = List.concat (map flatten exns); - -fun release_all results = - if List.all (fn Result _ => true | _ => false) results - then map (fn Result x => x) results - else raise EXCEPTIONS (flatten_list (List.mapPartial get_exn results)); - -fun release_first results = release_all results - handle EXCEPTIONS (exn :: _) => reraise exn; - -end; diff -r 63dd95e3b393 -r 3dcb46ae6185 src/Pure/ML-Systems/mosml.ML --- a/src/Pure/ML-Systems/mosml.ML Sat Dec 19 11:48:11 2009 +0100 +++ b/src/Pure/ML-Systems/mosml.ML Sat Dec 19 16:02:26 2009 +0100 @@ -41,7 +41,7 @@ fun reraise exn = raise exn; use "ML-Systems/unsynchronized.ML"; -use "ML-Systems/exn.ML"; +use "General/exn.ML"; use "ML-Systems/universal.ML"; use "ML-Systems/thread_dummy.ML"; use "ML-Systems/multithreading.ML"; diff -r 63dd95e3b393 -r 3dcb46ae6185 src/Pure/ML-Systems/polyml_common.ML --- a/src/Pure/ML-Systems/polyml_common.ML Sat Dec 19 11:48:11 2009 +0100 +++ b/src/Pure/ML-Systems/polyml_common.ML Sat Dec 19 16:02:26 2009 +0100 @@ -5,7 +5,7 @@ exception Interrupt = SML90.Interrupt; -use "ML-Systems/exn.ML"; +use "General/exn.ML"; use "ML-Systems/multithreading.ML"; use "ML-Systems/time_limit.ML"; use "ML-Systems/timing.ML"; diff -r 63dd95e3b393 -r 3dcb46ae6185 src/Pure/ML-Systems/smlnj.ML --- a/src/Pure/ML-Systems/smlnj.ML Sat Dec 19 11:48:11 2009 +0100 +++ b/src/Pure/ML-Systems/smlnj.ML Sat Dec 19 16:02:26 2009 +0100 @@ -9,7 +9,7 @@ use "ML-Systems/proper_int.ML"; use "ML-Systems/unsynchronized.ML"; use "ML-Systems/overloading_smlnj.ML"; -use "ML-Systems/exn.ML"; +use "General/exn.ML"; use "ML-Systems/universal.ML"; use "ML-Systems/thread_dummy.ML"; use "ML-Systems/multithreading.ML"; diff -r 63dd95e3b393 -r 3dcb46ae6185 src/Pure/library.scala --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/Pure/library.scala Sat Dec 19 16:02:26 2009 +0100 @@ -0,0 +1,24 @@ +/* Title: Pure/library.scala + Author: Makarius + +Basic library. +*/ + +package isabelle + +import java.lang.System + + +object Library +{ + /* timing */ + + def timeit[A](e: => A) = + { + val start = System.currentTimeMillis() + val result = Exn.capture(e) + val stop = System.currentTimeMillis() + System.err.println((stop - start) + "ms elapsed time") + Exn.release(result) + } +}