Pr.agda 1.8 KB

123456789101112131415161718192021222324252627282930313233343536373839404142434445464748495051525354555657585960616263646566676869707172737475767778798081828384858687888990919293
  1. module Pr where
  2. data FF : Set where
  3. magic : {X : Set} -> FF -> X
  4. magic ()
  5. record TT : Set where
  6. data Id {S : Set}(s : S) : S -> Set where
  7. refl : Id s s
  8. data Pr : Set1 where
  9. tt : Pr
  10. ff : Pr
  11. _/\_ : Pr -> Pr -> Pr
  12. all : (S : Set) -> (S -> Pr) -> Pr
  13. _eq_ : {S : Set} -> S -> S -> Pr
  14. record Sig (S : Set)(T : S -> Set) : Set where
  15. field
  16. fst : S
  17. snd : T fst
  18. open module Sig' {S : Set}{T : S -> Set} = Sig {S}{T} public
  19. _,_ : {S : Set}{T : S -> Set}(s : S) -> T s -> Sig S T
  20. s , t = record {fst = s ; snd = t}
  21. [|_|] : Pr -> Set
  22. [| tt |] = TT
  23. [| ff |] = FF
  24. [| P /\ Q |] = Sig [| P |] \_ -> [| Q |]
  25. [| all S P |] = (x : S) -> [| P x |]
  26. [| a eq b |] = Id a b
  27. _=>_ : Pr -> Pr -> Pr
  28. P => Q = all [| P |] \_ -> Q
  29. ∼ : Pr -> Pr
  30. ∼ P = P => ff
  31. data Decision (P : Pr) : Set where
  32. yes : [| P |] -> Decision P
  33. no : [| ∼ P |] -> Decision P
  34. data Bool : Set where
  35. true : Bool
  36. false : Bool
  37. So : Bool -> Pr
  38. So true = tt
  39. So false = ff
  40. not : Bool -> Bool
  41. not true = false
  42. not false = true
  43. so : (b : Bool) -> Decision (So b)
  44. so true = yes _
  45. so false = no magic
  46. potahto : (b : Bool) -> [| So (not b) => ∼ (So b) |]
  47. potahto true () _
  48. potahto false _ ()
  49. PEx : (P : Pr) -> ([| P |] -> Pr) -> Pr
  50. PEx P Q = P /\ all [| P |] Q
  51. Pow : Set -> Set1
  52. Pow X = X -> Pr
  53. _==>_ : {X : Set} -> Pow X -> Pow X -> Pr
  54. _==>_ {X} P Q = all X \x -> P x => Q x
  55. Decidable : {X : Set}(P : Pow X) -> Set
  56. Decidable {X} P = (x : X) -> Decision (P x)
  57. data _:-_ (S : Set)(P : Pow S) : Set where
  58. [_/_] : (s : S) -> [| P s |] -> S :- P
  59. wit : {S : Set}{P : S -> Pr} -> S :- P -> S
  60. wit [ s / p ] = s
  61. cert : {S : Set}{P : S -> Pr}(sp : S :- P) -> [| P (wit sp) |]
  62. cert [ s / p ] = p
  63. _??_ : {S : Set}{P : S -> Pr}
  64. (sp : S :- P){M : Set} ->
  65. ((s : S)(p : [| P s |]) -> M) ->
  66. M
  67. sp ?? m = m (wit sp) (cert sp)