src/HOL/ex/RPred.thy
author haftmann
Tue, 23 Nov 2010 23:11:06 +0100
changeset 40676 23904fa13e03
parent 36176 3fe7e97ccca8
child 42463 f270e3e18be5
permissions -rw-r--r--
merged
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
32308
c2b74affab85 imported patch generic compilation of predicate compiler with different monads
bulwahn
parents:
diff changeset
     1
theory RPred
c2b74affab85 imported patch generic compilation of predicate compiler with different monads
bulwahn
parents:
diff changeset
     2
imports Quickcheck Random Predicate
c2b74affab85 imported patch generic compilation of predicate compiler with different monads
bulwahn
parents:
diff changeset
     3
begin
c2b74affab85 imported patch generic compilation of predicate compiler with different monads
bulwahn
parents:
diff changeset
     4
33137
0d16c07f8d24 added option to generate random values to values command in the predicate compiler
bulwahn
parents: 32662
diff changeset
     5
types 'a "rpred" = "Random.seed \<Rightarrow> ('a Predicate.pred \<times> Random.seed)"
32308
c2b74affab85 imported patch generic compilation of predicate compiler with different monads
bulwahn
parents:
diff changeset
     6
c2b74affab85 imported patch generic compilation of predicate compiler with different monads
bulwahn
parents:
diff changeset
     7
section {* The RandomPred Monad *}
c2b74affab85 imported patch generic compilation of predicate compiler with different monads
bulwahn
parents:
diff changeset
     8
c2b74affab85 imported patch generic compilation of predicate compiler with different monads
bulwahn
parents:
diff changeset
     9
text {* A monad to combine the random state monad and the predicate monad *}
c2b74affab85 imported patch generic compilation of predicate compiler with different monads
bulwahn
parents:
diff changeset
    10
c2b74affab85 imported patch generic compilation of predicate compiler with different monads
bulwahn
parents:
diff changeset
    11
definition bot :: "'a rpred"
c2b74affab85 imported patch generic compilation of predicate compiler with different monads
bulwahn
parents:
diff changeset
    12
  where "bot = Pair (bot_class.bot)"
c2b74affab85 imported patch generic compilation of predicate compiler with different monads
bulwahn
parents:
diff changeset
    13
c2b74affab85 imported patch generic compilation of predicate compiler with different monads
bulwahn
parents:
diff changeset
    14
definition return :: "'a => 'a rpred"
c2b74affab85 imported patch generic compilation of predicate compiler with different monads
bulwahn
parents:
diff changeset
    15
  where "return x = Pair (Predicate.single x)"
c2b74affab85 imported patch generic compilation of predicate compiler with different monads
bulwahn
parents:
diff changeset
    16
32662
2faf1148c062 handling of definitions
bulwahn
parents: 32311
diff changeset
    17
definition bind :: "'a rpred \<Rightarrow> ('a \<Rightarrow> 'b rpred) \<Rightarrow> 'b rpred"
2faf1148c062 handling of definitions
bulwahn
parents: 32311
diff changeset
    18
(* (infixl "\<guillemotright>=" 60) *)
32308
c2b74affab85 imported patch generic compilation of predicate compiler with different monads
bulwahn
parents:
diff changeset
    19
  where "bind RP f =
c2b74affab85 imported patch generic compilation of predicate compiler with different monads
bulwahn
parents:
diff changeset
    20
  (\<lambda>s. let (P, s') = RP s;
c2b74affab85 imported patch generic compilation of predicate compiler with different monads
bulwahn
parents:
diff changeset
    21
        (s1, s2) = Random.split_seed s'
c2b74affab85 imported patch generic compilation of predicate compiler with different monads
bulwahn
parents:
diff changeset
    22
    in (Predicate.bind P (%a. fst (f a s1)), s2))"
c2b74affab85 imported patch generic compilation of predicate compiler with different monads
bulwahn
parents:
diff changeset
    23
32662
2faf1148c062 handling of definitions
bulwahn
parents: 32311
diff changeset
    24
definition supp :: "'a rpred \<Rightarrow> 'a rpred \<Rightarrow> 'a rpred"
2faf1148c062 handling of definitions
bulwahn
parents: 32311
diff changeset
    25
(* (infixl "\<squnion>" 80) *)
32308
c2b74affab85 imported patch generic compilation of predicate compiler with different monads
bulwahn
parents:
diff changeset
    26
where
c2b74affab85 imported patch generic compilation of predicate compiler with different monads
bulwahn
parents:
diff changeset
    27
  "supp RP1 RP2 = (\<lambda>s. let (P1, s') = RP1 s; (P2, s'') = RP2 s'
35028
108662d50512 more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
haftmann
parents: 33142
diff changeset
    28
  in (semilattice_sup_class.sup P1 P2, s''))"
32308
c2b74affab85 imported patch generic compilation of predicate compiler with different monads
bulwahn
parents:
diff changeset
    29
c2b74affab85 imported patch generic compilation of predicate compiler with different monads
bulwahn
parents:
diff changeset
    30
definition if_rpred :: "bool \<Rightarrow> unit rpred"
c2b74affab85 imported patch generic compilation of predicate compiler with different monads
bulwahn
parents:
diff changeset
    31
where
c2b74affab85 imported patch generic compilation of predicate compiler with different monads
bulwahn
parents:
diff changeset
    32
  "if_rpred b = (if b then return () else bot)"
c2b74affab85 imported patch generic compilation of predicate compiler with different monads
bulwahn
parents:
diff changeset
    33
c2b74affab85 imported patch generic compilation of predicate compiler with different monads
bulwahn
parents:
diff changeset
    34
(* Missing a good definition for negation: not_rpred *)
c2b74affab85 imported patch generic compilation of predicate compiler with different monads
bulwahn
parents:
diff changeset
    35
33138
e2e23987c59a reinvestigating the compilation of the random computation in the predicate compiler
bulwahn
parents: 33137
diff changeset
    36
definition not_rpred :: "unit rpred \<Rightarrow> unit rpred"
32308
c2b74affab85 imported patch generic compilation of predicate compiler with different monads
bulwahn
parents:
diff changeset
    37
where
33138
e2e23987c59a reinvestigating the compilation of the random computation in the predicate compiler
bulwahn
parents: 33137
diff changeset
    38
  "not_rpred P = (\<lambda>s. let (P', s') = P s in if Predicate.eval P' () then (Orderings.bot, s') else (Predicate.single (), s'))"
32308
c2b74affab85 imported patch generic compilation of predicate compiler with different monads
bulwahn
parents:
diff changeset
    39
c2b74affab85 imported patch generic compilation of predicate compiler with different monads
bulwahn
parents:
diff changeset
    40
definition lift_pred :: "'a Predicate.pred \<Rightarrow> 'a rpred"
c2b74affab85 imported patch generic compilation of predicate compiler with different monads
bulwahn
parents:
diff changeset
    41
  where
c2b74affab85 imported patch generic compilation of predicate compiler with different monads
bulwahn
parents:
diff changeset
    42
  "lift_pred = Pair"
c2b74affab85 imported patch generic compilation of predicate compiler with different monads
bulwahn
parents:
diff changeset
    43
32311
50f953140636 adapted predicate compiler for quickcheck generators; added compilation of depth-limited search functions for predicates
bulwahn
parents: 32308
diff changeset
    44
definition lift_random :: "(Random.seed \<Rightarrow> ('a \<times> (unit \<Rightarrow> term)) \<times> Random.seed) \<Rightarrow> 'a rpred"
50f953140636 adapted predicate compiler for quickcheck generators; added compilation of depth-limited search functions for predicates
bulwahn
parents: 32308
diff changeset
    45
  where "lift_random g = scomp g (Pair o (Predicate.single o fst))"
32308
c2b74affab85 imported patch generic compilation of predicate compiler with different monads
bulwahn
parents:
diff changeset
    46
33137
0d16c07f8d24 added option to generate random values to values command in the predicate compiler
bulwahn
parents: 32662
diff changeset
    47
definition map :: "('a \<Rightarrow> 'b) \<Rightarrow> ('a rpred \<Rightarrow> 'b rpred)"
33142
edab304696ec adapted parser for options in the predicate compiler
bulwahn
parents: 33138
diff changeset
    48
  where "map f P = bind P (return o f)"
32662
2faf1148c062 handling of definitions
bulwahn
parents: 32311
diff changeset
    49
36176
3fe7e97ccca8 replaced generic 'hide' command by more conventional 'hide_class', 'hide_type', 'hide_const', 'hide_fact' -- frees some popular keywords;
wenzelm
parents: 35028
diff changeset
    50
hide_const (open) return bind supp map
32308
c2b74affab85 imported patch generic compilation of predicate compiler with different monads
bulwahn
parents:
diff changeset
    51
  
c2b74affab85 imported patch generic compilation of predicate compiler with different monads
bulwahn
parents:
diff changeset
    52
end