1493 @{define_ML the: "'a option -> 'a"} \\ |
1493 @{define_ML the: "'a option -> 'a"} \\ |
1494 @{define_ML these: "'a list option -> 'a list"} \\ |
1494 @{define_ML these: "'a list option -> 'a list"} \\ |
1495 @{define_ML the_list: "'a option -> 'a list"} \\ |
1495 @{define_ML the_list: "'a option -> 'a list"} \\ |
1496 @{define_ML the_default: "'a -> 'a option -> 'a"} \\ |
1496 @{define_ML the_default: "'a -> 'a option -> 'a"} \\ |
1497 \end{mldecls} |
1497 \end{mldecls} |
|
1498 |
|
1499 \begin{matharray}{rcl} |
|
1500 @{ML_antiquotation_def "if_none"} & : & \<open>ML_antiquotation\<close> \\ |
|
1501 \end{matharray} |
|
1502 |
|
1503 \<^rail>\<open>@@{ML_antiquotation if_none} embedded\<close> |
1498 \<close> |
1504 \<close> |
1499 |
1505 |
1500 text \<open> |
1506 text \<open> |
1501 Apart from \<^ML>\<open>Option.map\<close> most other operations defined in structure |
1507 Apart from \<^ML>\<open>Option.map\<close> most other operations defined in structure |
1502 \<^ML_structure>\<open>Option\<close> are alien to Isabelle/ML and never used. The |
1508 \<^ML_structure>\<open>Option\<close> are alien to Isabelle/ML and never used. The |
1503 operations shown above are defined in \<^file>\<open>~~/src/Pure/General/basics.ML\<close>. |
1509 operations shown above are defined in \<^file>\<open>~~/src/Pure/General/basics.ML\<close>. |
|
1510 |
|
1511 Note that the function \<^ML>\<open>the_default\<close> is strict in all of its |
|
1512 arguments, the default value is evaluated beforehand, even if not required |
|
1513 later. In contrast, the antiquotation @{ML_antiquotation "if_none"} is |
|
1514 non-strict: the given expression is only evaluated for an application to |
|
1515 \<^ML>\<open>NONE\<close>. This allows to work with exceptions like this: |
|
1516 \<close> |
|
1517 |
|
1518 ML \<open> |
|
1519 fun div_total x y = |
|
1520 \<^try>\<open>x div y\<close> |> the_default 0; |
|
1521 |
|
1522 fun div_error x y = |
|
1523 \<^try>\<open>x div y\<close> |> \<^if_none>\<open>error "Division by zero"\<close>; |
|
1524 \<close> |
|
1525 |
|
1526 text \<open> |
|
1527 Of course, it is also possible to handle exceptions directly, without an |
|
1528 intermediate option construction: |
|
1529 \<close> |
|
1530 |
|
1531 ML \<open> |
|
1532 fun div_total x y = |
|
1533 x div y handle Div => 0; |
|
1534 |
|
1535 fun div_error x y = |
|
1536 x div y handle Div => error "Division by zero"; |
|
1537 \<close> |
|
1538 |
|
1539 text \<open> |
|
1540 The first form works better in longer chains of functional composition, with |
|
1541 combinators like \<^ML>\<open>|>\<close> or \<^ML>\<open>#>\<close> or \<^ML>\<open>o\<close>. The second form is more |
|
1542 adequate in elementary expressions: there is no need to pretend that |
|
1543 Isabelle/ML is actually a version of Haskell. |
1504 \<close> |
1544 \<close> |
1505 |
1545 |
1506 |
1546 |
1507 subsection \<open>Lists\<close> |
1547 subsection \<open>Lists\<close> |
1508 |
1548 |