added Isar/value_parse.ML: Outer syntax parsers for basic ML values.
--- a/src/Pure/IsaMakefile Fri Jan 02 15:44:32 2009 +0100
+++ b/src/Pure/IsaMakefile Fri Jan 02 15:44:32 2009 +0100
@@ -47,15 +47,15 @@
Isar/proof_display.ML Isar/proof_node.ML Isar/rule_cases.ML \
Isar/rule_insts.ML Isar/session.ML Isar/skip_proof.ML \
Isar/spec_parse.ML Isar/specification.ML Isar/subclass.ML \
- Isar/theory_target.ML Isar/toplevel.ML ML-Systems/alice.ML \
- ML-Systems/exn.ML ML-Systems/install_pp_polyml.ML \
- ML-Systems/ml_name_space.ML ML-Systems/multithreading.ML \
- ML-Systems/mosml.ML ML-Systems/multithreading_polyml.ML \
- ML-Systems/overloading_smlnj.ML ML-Systems/polyml-4.1.3.ML \
- ML-Systems/polyml-4.1.4.ML ML-Systems/polyml-4.2.0.ML \
- ML-Systems/polyml-5.0.ML ML-Systems/polyml-5.1.ML \
- ML-Systems/polyml_common.ML ML-Systems/polyml.ML \
- ML-Systems/polyml_old_compiler4.ML \
+ Isar/theory_target.ML Isar/toplevel.ML Isar/value_parse.ML \
+ ML-Systems/alice.ML ML-Systems/exn.ML \
+ ML-Systems/install_pp_polyml.ML ML-Systems/ml_name_space.ML \
+ ML-Systems/multithreading.ML ML-Systems/mosml.ML \
+ ML-Systems/multithreading_polyml.ML ML-Systems/overloading_smlnj.ML \
+ ML-Systems/polyml-4.1.3.ML ML-Systems/polyml-4.1.4.ML \
+ ML-Systems/polyml-4.2.0.ML ML-Systems/polyml-5.0.ML \
+ ML-Systems/polyml-5.1.ML ML-Systems/polyml_common.ML \
+ ML-Systems/polyml.ML ML-Systems/polyml_old_compiler4.ML \
ML-Systems/polyml_old_compiler5.ML ML-Systems/proper_int.ML \
ML-Systems/smlnj.ML ML-Systems/system_shell.ML \
ML-Systems/time_limit.ML ML-Systems/thread_dummy.ML \
--- a/src/Pure/Isar/ROOT.ML Fri Jan 02 15:44:32 2009 +0100
+++ b/src/Pure/Isar/ROOT.ML Fri Jan 02 15:44:32 2009 +0100
@@ -1,5 +1,4 @@
(* Title: Pure/Isar/ROOT.ML
- ID: $Id$
Author: Markus Wenzel, TU Muenchen
Isar -- Intelligible Semi-Automated Reasoning for Isabelle.
@@ -21,8 +20,9 @@
(*outer syntax*)
use "outer_lex.ML";
+use "outer_keyword.ML";
use "outer_parse.ML";
-use "outer_keyword.ML";
+use "value_parse.ML";
use "args.ML";
use "antiquote.ML";
use "../ML/ml_context.ML";
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/Pure/Isar/value_parse.ML Fri Jan 02 15:44:32 2009 +0100
@@ -0,0 +1,45 @@
+(* Title: Pure/Isar/value_parse.ML
+ Author: Makarius
+
+Outer syntax parsers for basic ML values.
+*)
+
+signature VALUE_PARSE =
+sig
+ type 'a parser = 'a OuterParse.parser
+ val comma: 'a parser -> 'a parser
+ val equal: 'a parser -> 'a parser
+ val parens: 'a parser -> 'a parser
+ val pair: 'a parser -> 'b parser -> ('a * 'b) parser
+ val triple: 'a parser -> 'b parser -> 'c parser -> ('a * 'b * 'c) parser
+ val properties: Properties.T parser
+end;
+
+structure ValueParse: VALUE_PARSE =
+struct
+
+structure P = OuterParse;
+type 'a parser = 'a P.parser;
+
+
+(* syntax utilities *)
+
+fun comma p = P.$$$ "," |-- P.!!! p;
+fun equal p = P.$$$ "=" |-- P.!!! p;
+fun parens p = P.$$$ "(" |-- P.!!! (p --| P.$$$ ")");
+
+
+(* tuples *)
+
+val unit = parens (Scan.succeed ());
+fun pair p1 p2 = parens (p1 -- comma p2);
+fun triple p1 p2 p3 = parens (p1 -- comma p2 -- comma p3) >> P.triple1;
+
+
+(* lists *)
+
+fun list p = parens (P.enum "," p);
+val properties = list (P.string -- equal P.string);
+
+end;
+