# HG changeset patch # User wenzelm # Date 1682189705 -7200 # Node ID ee9785abbcd6a156f1c4a3c2a2ea0b5ebee809fe # Parent 9c5e8460df059c65bbe152eb33c3cba2efcf8d3d provide ML antiquotation "if_none": non-strict version of "the_default"; diff -r 9c5e8460df05 -r ee9785abbcd6 NEWS --- a/NEWS Sat Apr 22 10:22:41 2023 +0200 +++ b/NEWS Sat Apr 22 20:55:05 2023 +0200 @@ -283,6 +283,10 @@ *** ML *** +* ML antiquotation \<^if_none>\expr\ inlines a function (fn SOME x => x +| NONE => expr); this is a non-strict version of the regular function +"the_default". Both are particularly useful with the |> combinator. + * Improved implementations and signatures of functor Table() and corresponding functor Set(). diff -r 9c5e8460df05 -r ee9785abbcd6 etc/symbols --- a/etc/symbols Sat Apr 22 10:22:41 2023 +0200 +++ b/etc/symbols Sat Apr 22 20:55:05 2023 +0200 @@ -506,6 +506,7 @@ \<^if_macos> argument: cartouche \<^if_windows> argument: cartouche \<^if_unix> argument: cartouche +\<^if_none> argument: cartouche \<^cite> argument: cartouche \<^nocite> argument: cartouche \<^citet> argument: cartouche diff -r 9c5e8460df05 -r ee9785abbcd6 lib/texinputs/isabellesym.sty --- a/lib/texinputs/isabellesym.sty Sat Apr 22 10:22:41 2023 +0200 +++ b/lib/texinputs/isabellesym.sty Sat Apr 22 20:55:05 2023 +0200 @@ -495,6 +495,7 @@ \newcommand{\isactrlifUNDERSCOREmacos}{\isakeywordcontrol{if{\isacharunderscore}macos}} \newcommand{\isactrlifUNDERSCOREwindows}{\isakeywordcontrol{if{\isacharunderscore}windows}} \newcommand{\isactrlifUNDERSCOREunix}{\isakeywordcontrol{if{\isacharunderscore}unix}} +\newcommand{\isactrlifUNDERSCOREnone}{\isakeywordcontrol{if{\isacharunderscore}none}} \newcommand{\isactrlcite}{\isakeywordcontrol{cite}} \newcommand{\isactrlnocite}{\isakeywordcontrol{nocite}} diff -r 9c5e8460df05 -r ee9785abbcd6 src/Doc/Implementation/ML.thy --- a/src/Doc/Implementation/ML.thy Sat Apr 22 10:22:41 2023 +0200 +++ b/src/Doc/Implementation/ML.thy Sat Apr 22 20:55:05 2023 +0200 @@ -1495,12 +1495,52 @@ @{define_ML the_list: "'a option -> 'a list"} \\ @{define_ML the_default: "'a -> 'a option -> 'a"} \\ \end{mldecls} + + \begin{matharray}{rcl} + @{ML_antiquotation_def "if_none"} & : & \ML_antiquotation\ \\ + \end{matharray} + + \<^rail>\@@{ML_antiquotation if_none} embedded\ \ text \ Apart from \<^ML>\Option.map\ most other operations defined in structure \<^ML_structure>\Option\ are alien to Isabelle/ML and never used. The operations shown above are defined in \<^file>\~~/src/Pure/General/basics.ML\. + + Note that the function \<^ML>\the_default\ is strict in all of its + arguments, the default value is evaluated beforehand, even if not required + later. In contrast, the antiquotation @{ML_antiquotation "if_none"} is + non-strict: the given expression is only evaluated for an application to + \<^ML>\NONE\. This allows to work with exceptions like this: +\ + +ML \ + fun div_total x y = + \<^try>\x div y\ |> the_default 0; + + fun div_error x y = + \<^try>\x div y\ |> \<^if_none>\error "Division by zero"\; +\ + +text \ + Of course, it is also possible to handle exceptions directly, without an + intermediate option construction: +\ + +ML \ + fun div_total x y = + x div y handle Div => 0; + + fun div_error x y = + x div y handle Div => error "Division by zero"; +\ + +text \ + The first form works better in longer chains of functional composition, with + combinators like \<^ML>\|>\ or \<^ML>\#>\ or \<^ML>\o\. The second form is more + adequate in elementary expressions: there is no need to pretend that + Isabelle/ML is actually a version of Haskell. \ diff -r 9c5e8460df05 -r ee9785abbcd6 src/Pure/ML/ml_antiquotations.ML --- a/src/Pure/ML/ml_antiquotations.ML Sat Apr 22 10:22:41 2023 +0200 +++ b/src/Pure/ML/ml_antiquotations.ML Sat Apr 22 20:55:05 2023 +0200 @@ -340,11 +340,24 @@ in end; -(* special forms *) +(* special forms for option type *) val _ = Theory.setup (ML_Antiquotation.special_form \<^binding>\try\ "() |> Basics.try" #> - ML_Antiquotation.special_form \<^binding>\can\ "() |> Basics.can"); + ML_Antiquotation.special_form \<^binding>\can\ "() |> Basics.can" #> + ML_Context.add_antiquotation \<^binding>\if_none\ true + (fn _ => fn src => fn ctxt => + let + val (s, _) = Token.syntax (Scan.lift Parse.embedded_input) src ctxt; + val tokenize = ML_Lex.tokenize_no_range; + + val (decl, ctxt') = ML_Context.expand_antiquotes (ML_Lex.read_source s) ctxt; + fun decl' ctxt'' = + let + val (ml_env, ml_body) = decl ctxt''; + val ml_body' = tokenize "(fn SOME x => x | NONE => " @ ml_body @ tokenize ")"; + in (ml_env, ml_body') end; + in (decl', ctxt') end)); (* basic combinators *)