diff -r c8a2755bf220 -r ff784d5a5bfb src/HOL/Imperative_HOL/Overview.thy --- a/src/HOL/Imperative_HOL/Overview.thy Sat Jan 05 17:00:43 2019 +0100 +++ b/src/HOL/Imperative_HOL/Overview.thy Sat Jan 05 17:24:33 2019 +0100 @@ -33,22 +33,20 @@ section \A polymorphic heap inside a monad\ text \ - Heaps (@{type heap}) can be populated by values of class @{class - heap}; HOL's default types are already instantiated to class @{class - heap}. Class @{class heap} is a subclass of @{class countable}; see - theory \Countable\ for ways to instantiate types as @{class countable}. + Heaps (\<^type>\heap\) can be populated by values of class \<^class>\heap\; HOL's default types are already instantiated to class \<^class>\heap\. Class \<^class>\heap\ is a subclass of \<^class>\countable\; see + theory \Countable\ for ways to instantiate types as \<^class>\countable\. - The heap is wrapped up in a monad @{typ "'a Heap"} by means of the + The heap is wrapped up in a monad \<^typ>\'a Heap\ by means of the following specification: \begin{quote} - @{datatype Heap} + \<^datatype>\Heap\ \end{quote} Unwrapping of this monad type happens through \begin{quote} - @{term_type execute} \\ + \<^term_type>\execute\ \\ @{thm execute.simps [no_vars]} \end{quote} @@ -66,31 +64,30 @@ Monadic expression involve the usual combinators: \begin{quote} - @{term_type return} \\ - @{term_type bind} \\ - @{term_type raise} + \<^term_type>\return\ \\ + \<^term_type>\bind\ \\ + \<^term_type>\raise\ \end{quote} - This is also associated with nice monad do-syntax. The @{typ - string} argument to @{const raise} is just a codified comment. + This is also associated with nice monad do-syntax. The \<^typ>\string\ argument to \<^const>\raise\ is just a codified comment. Among a couple of generic combinators the following is helpful for establishing invariants: \begin{quote} - @{term_type assert} \\ + \<^term_type>\assert\ \\ @{thm assert_def [no_vars]} \end{quote} \ -section \Relational reasoning about @{type Heap} expressions\ +section \Relational reasoning about \<^type>\Heap\ expressions\ text \ To establish correctness of imperative programs, predicate \begin{quote} - @{term_type effect} + \<^term_type>\effect\ \end{quote} provides a simple relational calculus. Primitive rules are \effectI\ and \effectE\, rules appropriate for reasoning about @@ -101,13 +98,13 @@ on the heap at all; reasoning then can be easier using predicate \begin{quote} - @{term_type success} + \<^term_type>\success\ \end{quote} - Introduction rules for @{const success} are available in the + Introduction rules for \<^const>\success\ are available in the \success_intro\ fact collection. - @{const execute}, @{const effect}, @{const success} and @{const bind} + \<^const>\execute\, \<^const>\effect\, \<^const>\success\ and \<^const>\bind\ are related by rules \execute_bind_success\, \success_bind_executeI\, \success_bind_effectI\, \effect_bindI\, \effect_bindE\ and \execute_bind_eq_SomeI\. \ @@ -141,27 +138,27 @@ Heap operations: \begin{quote} - @{term_type Array.alloc} \\ - @{term_type Array.present} \\ - @{term_type Array.get} \\ - @{term_type Array.set} \\ - @{term_type Array.length} \\ - @{term_type Array.update} \\ - @{term_type Array.noteq} + \<^term_type>\Array.alloc\ \\ + \<^term_type>\Array.present\ \\ + \<^term_type>\Array.get\ \\ + \<^term_type>\Array.set\ \\ + \<^term_type>\Array.length\ \\ + \<^term_type>\Array.update\ \\ + \<^term_type>\Array.noteq\ \end{quote} Monad operations: \begin{quote} - @{term_type Array.new} \\ - @{term_type Array.of_list} \\ - @{term_type Array.make} \\ - @{term_type Array.len} \\ - @{term_type Array.nth} \\ - @{term_type Array.upd} \\ - @{term_type Array.map_entry} \\ - @{term_type Array.swap} \\ - @{term_type Array.freeze} + \<^term_type>\Array.new\ \\ + \<^term_type>\Array.of_list\ \\ + \<^term_type>\Array.make\ \\ + \<^term_type>\Array.len\ \\ + \<^term_type>\Array.nth\ \\ + \<^term_type>\Array.upd\ \\ + \<^term_type>\Array.map_entry\ \\ + \<^term_type>\Array.swap\ \\ + \<^term_type>\Array.freeze\ \end{quote} \ @@ -171,20 +168,20 @@ Heap operations: \begin{quote} - @{term_type Ref.alloc} \\ - @{term_type Ref.present} \\ - @{term_type Ref.get} \\ - @{term_type Ref.set} \\ - @{term_type Ref.noteq} + \<^term_type>\Ref.alloc\ \\ + \<^term_type>\Ref.present\ \\ + \<^term_type>\Ref.get\ \\ + \<^term_type>\Ref.set\ \\ + \<^term_type>\Ref.noteq\ \end{quote} Monad operations: \begin{quote} - @{term_type Ref.ref} \\ - @{term_type Ref.lookup} \\ - @{term_type Ref.update} \\ - @{term_type Ref.change} + \<^term_type>\Ref.ref\ \\ + \<^term_type>\Ref.lookup\ \\ + \<^term_type>\Ref.update\ \\ + \<^term_type>\Ref.change\ \end{quote} \