src/HOL/Probability/PMF_Impl.thy
changeset 72607 feebdaa346e5
parent 72581 de581f98a3a1
child 81113 6fefd6c602fa
--- a/src/HOL/Probability/PMF_Impl.thy	Sun Nov 15 07:17:05 2020 +0000
+++ b/src/HOL/Probability/PMF_Impl.thy	Sun Nov 15 07:17:06 2020 +0000
@@ -530,22 +530,22 @@
 definition single :: "'a \<Rightarrow> 'a multiset" where
 "single s = {#s#}"
 
-definition (in term_syntax)
-  pmfify :: "('a::typerep multiset \<times> (unit \<Rightarrow> Code_Evaluation.term)) \<Rightarrow>
-             'a \<times> (unit \<Rightarrow> Code_Evaluation.term) \<Rightarrow>
-             'a pmf \<times> (unit \<Rightarrow> Code_Evaluation.term)" where
+instantiation pmf :: (random) random
+begin
+
+context
+  includes state_combinator_syntax term_syntax
+begin
+
+definition
+  pmfify :: "('b::typerep multiset \<times> (unit \<Rightarrow> Code_Evaluation.term)) \<Rightarrow>
+             'b \<times> (unit \<Rightarrow> Code_Evaluation.term) \<Rightarrow>
+             'b pmf \<times> (unit \<Rightarrow> Code_Evaluation.term)" where
   [code_unfold]: "pmfify A x =  
     Code_Evaluation.valtermify pmf_of_multiset {\<cdot>} 
       (Code_Evaluation.valtermify (+) {\<cdot>} A {\<cdot>} 
        (Code_Evaluation.valtermify single {\<cdot>} x))"
 
-instantiation pmf :: (random) random
-begin
-
-context
-  includes state_combinator_syntax
-begin
-
 definition
   "Quickcheck_Random.random i = 
      Quickcheck_Random.random i \<circ>\<rightarrow> (\<lambda>A.