528 end |
528 end |
529 |
529 |
530 definition single :: "'a \<Rightarrow> 'a multiset" where |
530 definition single :: "'a \<Rightarrow> 'a multiset" where |
531 "single s = {#s#}" |
531 "single s = {#s#}" |
532 |
532 |
533 definition (in term_syntax) |
533 instantiation pmf :: (random) random |
534 pmfify :: "('a::typerep multiset \<times> (unit \<Rightarrow> Code_Evaluation.term)) \<Rightarrow> |
534 begin |
535 'a \<times> (unit \<Rightarrow> Code_Evaluation.term) \<Rightarrow> |
535 |
536 'a pmf \<times> (unit \<Rightarrow> Code_Evaluation.term)" where |
536 context |
|
537 includes state_combinator_syntax term_syntax |
|
538 begin |
|
539 |
|
540 definition |
|
541 pmfify :: "('b::typerep multiset \<times> (unit \<Rightarrow> Code_Evaluation.term)) \<Rightarrow> |
|
542 'b \<times> (unit \<Rightarrow> Code_Evaluation.term) \<Rightarrow> |
|
543 'b pmf \<times> (unit \<Rightarrow> Code_Evaluation.term)" where |
537 [code_unfold]: "pmfify A x = |
544 [code_unfold]: "pmfify A x = |
538 Code_Evaluation.valtermify pmf_of_multiset {\<cdot>} |
545 Code_Evaluation.valtermify pmf_of_multiset {\<cdot>} |
539 (Code_Evaluation.valtermify (+) {\<cdot>} A {\<cdot>} |
546 (Code_Evaluation.valtermify (+) {\<cdot>} A {\<cdot>} |
540 (Code_Evaluation.valtermify single {\<cdot>} x))" |
547 (Code_Evaluation.valtermify single {\<cdot>} x))" |
541 |
|
542 instantiation pmf :: (random) random |
|
543 begin |
|
544 |
|
545 context |
|
546 includes state_combinator_syntax |
|
547 begin |
|
548 |
548 |
549 definition |
549 definition |
550 "Quickcheck_Random.random i = |
550 "Quickcheck_Random.random i = |
551 Quickcheck_Random.random i \<circ>\<rightarrow> (\<lambda>A. |
551 Quickcheck_Random.random i \<circ>\<rightarrow> (\<lambda>A. |
552 Quickcheck_Random.random i \<circ>\<rightarrow> (\<lambda>x. Pair (pmfify A x)))" |
552 Quickcheck_Random.random i \<circ>\<rightarrow> (\<lambda>x. Pair (pmfify A x)))" |