diff -r c09598a97436 -r d8d85a8172b5 src/HOL/Random.thy --- a/src/HOL/Random.thy Sat Jul 18 21:44:18 2015 +0200 +++ b/src/HOL/Random.thy Sat Jul 18 22:58:50 2015 +0200 @@ -1,7 +1,7 @@ (* Author: Florian Haftmann, TU Muenchen *) -section {* A HOL random engine *} +section \A HOL random engine\ theory Random imports List Groups_List @@ -11,7 +11,7 @@ notation scomp (infixl "\\" 60) -subsection {* Auxiliary functions *} +subsection \Auxiliary functions\ fun log :: "natural \ natural \ natural" where "log b i = (if b \ 1 \ i < b then 1 else 1 + log b (i div b))" @@ -23,7 +23,7 @@ "minus_shift r k l = (if k < l then r + k - l else k - l)" -subsection {* Random seeds *} +subsection \Random seeds\ type_synonym seed = "natural \ natural" @@ -45,7 +45,7 @@ in ((v'', w'), (v', w'')))" -subsection {* Base selectors *} +subsection \Base selectors\ fun iterate :: "natural \ ('b \ 'a \ 'b \ 'a) \ 'b \ 'a \ 'b \ 'a" where "iterate k f x = (if k = 0 then Pair x else f x \\ iterate (k - 1) f)" @@ -136,12 +136,12 @@ qed -subsection {* @{text ML} interface *} +subsection \@{text ML} interface\ code_reflect Random_Engine functions range select select_weight -ML {* +ML \ structure Random_Engine = struct @@ -177,7 +177,7 @@ end; end; -*} +\ hide_type (open) seed hide_const (open) inc_shift minus_shift log "next" split_seed