Probabilistic Functions and Cryptographic Oracles in Higher Order Logic

Size: px
Start display at page:

Download "Probabilistic Functions and Cryptographic Oracles in Higher Order Logic"

Transcription

1 Probabilistic Functions and Cryptographic Oracles in Higher Order Logic Andreas Lochbihler March 2, 2016 Contents 1 Auxiliary material Countable Extended naturals Indicator function Integrals Measures Sequence space PMF Streams Transfer rules Algebraic stuff Option Orders on option filter for option map for the combination of set and option join on options Zip on options Predicator for function relator with relation op = on the domain A relator for sets that treats sets like predicates List of a given length Corecursor for products If Monomorphic monads and monad transformers Plain monads Identity monad Monads with failure The option monad The set monad Monad homomorphism

2 3 Monad transformers The state monad transformer Transformation of homomorphisms The reader/environment monad transformer Transformation of homomorphisms Subprobability mass functions Support Functorial structure Monad operations Return Bind Relator From a pmf to a spmf Weight of a subprobability From density to spmfs Ordering on spmfs Restrictions on spmfs Integral over spmf Non-negative integral Lebesgue integral Bounded functions Subprobability distributions Losslessness Scaling Conditional spmfs Product spmf Assertions Try Embedding of a option into a spmf The IO monad The real world Setup for lifting and transfer Code generator setup Semantics of the IO monad The resumption monad setup for partial-function Execution resumption Code generator setup Setup for lifting and transfer Single-step generative

3 7 Generative probabilistic values Type definition Simple, derived operations Monad structure Embedding a spmf as a monad Embedding a option as a monad Embedding resumptions Assertions Executing gpv: possibilistic trace semantics Order for ( a, out, in) gpv Bounds on interaction Typing Interface between gpvs and rpvs Type judgements Sub-gpvs Losslessness Inlining Running GPVs Bisimulation for oracles Oracle combinators Combining oracles Negligibility Applicative Functors The basic operations of applicative functors Applicative functors with combinatorial structure The environment functor lifting predicates to the environment functor Cyclic groups Bitstrings Examples The IND-CPA security game for public-key encryption with oracle access The IND-CPA security game (public key, single instance) The IND-CPA security game The DDH security game The LCDH security game Random function construction Pseudo-random functions The bernoulli distribution built from coin flips

4 13.9 Elgamal encryption scheme Formalisation of Hashed Elgamal in the Random Oracle Model IND-CPA from PRF

5 theory Probability-More imports /src/hol/probability/probability begin 1 Auxiliary material lemma csup-singleton [simp]: (SUP x:{x}. f x :: - :: conditionally-complete-lattice) = f x unfolding SUP-def by simp Countable lemma countable-lfp: assumes step: Y. countable Y = countable (F Y ) and cont: Order-Continuity.sup-continuous F shows countable (lfp F ) by(subst sup-continuous-lfp[of cont])(simp add: countable-funpow[of step]) lemma countable-lfp-apply: assumes step: Y x. ( x. countable (Y x)) = countable (F Y x) and cont: Order-Continuity.sup-continuous F shows countable (lfp F x) proof { fix n have x. countable ((F ˆˆ n) bot x) by(induct n)(auto intro: step) } thus?thesis using cont by(simp add: sup-continuous-lfp) Extended naturals lemma idiff-enat-eq-enat-iff : x enat n = enat m ( k. x = enat k k n = m) by(cases x) simp-all Indicator function lemma indicator-single-some: indicator {Some x} (Some y) = indicator {x} y by(simp split: split-indicator) Integrals lemma (in finite-measure) nn-integral-indicator-neq-infty: f A sets M = ( + x. indicator A (f x) M ) unfolding ereal-indicator[symmetric] apply(rule integrabled) apply(rule integrable-const-bound[where B =1 ]) apply(simp-all add: indicator-vimage[symmetric]) done lemma nn-integral-indicator-map: 5

6 assumes [measurable]: f measurable M N {x space N. P x} sets N shows ( + x. indicator {x space N. P x} (f x) M ) = emeasure M {x space M. P (f x)} using assms(1 )[THEN measurable-space] by (subst nn-integral-indicator[symmetric]) (auto intro!: nn-integral-cong split: split-indicator simp del: nn-integral-indicator) Measures declare sets-restrict-space-count-space [measurable-cong] lemma (in sigma-algebra) sets-collect-countable-ex1 : ( i :: i :: countable. {x Ω. P i x} M ) = {x Ω.!i. P i x} M using sets-collect-countable-ex1 [of UNIV :: i set] by simp lemma pred-countable-ex1 [measurable]: ( i :: - :: countable. Measurable.pred M (λx. P i x)) = Measurable.pred M (λx.!i. P i x) unfolding pred-def by(rule sets.sets-collect-countable-ex1 ) lemma measurable-snd-count-space [measurable]: A B = snd measurable (M1 M count-space A) (count-space B) by(auto simp add: measurable-def space-pair-measure snd-vimage-eq-times times-int-times) Sequence space lemma (in sequence-space) nn-integral-split: assumes f [measurable]: f borel-measurable S shows ( + ω. f ω S) = ( + ω. ( + ω. f (comb-seq i ω ω ) S) S) by (subst PiM-comb-seq[symmetric, where i =i]) (simp add: nn-integral-distr P.nn-integral-fst[symmetric]) lemma (in sequence-space) prob-collect-split: assumes f [measurable]: {x space S. P x} sets S shows P(x in S. P x) = ( + x. P(x in S. P (comb-seq i x x )) S) proof have P(x in S. P x) = ( + x. ( + x. indicator {x space S. P x} (comb-seq i x x ) S) S) using nn-integral-split[of indicator {x space S. P x}] by (auto simp: emeasure-eq-measure) also have... = ( + x. P(x in S. P (comb-seq i x x )) S) by (intro nn-integral-cong) (auto simp: emeasure-eq-measure nn-integral-indicator-map) finally show?thesis PMF adhoc-overloading Monad-Syntax.bind bind-pmf lemma measure-map-pmf-conv-distr: measure-pmf (map-pmf f p) = distr (measure-pmf p) (count-space UNIV ) f 6

7 by(fact map-pmf-rep-eq) Streams primrec sprefix :: a list a stream bool where sprefix-nil: sprefix [] ys = True sprefix-cons: sprefix (x # xs) ys x = shd ys sprefix xs (stl ys) lemma sprefix-append: sprefix ys) zs sprefix xs zs sprefix ys (sdrop (length xs) zs) by(induct xs arbitrary: zs) simp-all lemma sprefix-stake-same [simp]: sprefix (stake n xs) xs by(induct n arbitrary: xs) simp-all lemma sprefix-same-imp-eq: assumes sprefix xs ys sprefix xs ys and length xs = length xs shows xs = xs using assms(3,1,2 ) by(induct arbitrary: ys rule: list-induct2 ) auto lemma sprefix-shift-same [simp]: sprefix xs ys) by(induct xs) simp-all lemma snth-sdrop: sdrop n xs!! m = xs!! (n + m) by(induct n arbitrary: xs) simp-all lemma sprefix-shift [simp]: length xs length ys = sprefix xs zs) prefixeq xs ys by(induct xs arbitrary: ys)(simp, case-tac ys, auto) lemma prefixeq-stake2 [simp]: prefixeq xs (stake n ys) length xs n sprefix xs ys proof(induct xs arbitrary: n ys) case (Cons x xs) thus?case by(cases ys n rule: stream.exhaust[case-product nat.exhaust]) auto simp lemma sdrop-shift: sdrop n ys) = drop n sdrop (n length xs) ys by(induct xs arbitrary: n)(simp, case-tac n, simp-all) end theory Misc imports Complex-Main /src/hol/eisbach/eisbach /src/hol/library/rewrite /src/hol/library/simps-case-conv begin 7

8 lemma Least-le-Least: fixes x :: a :: wellorder assumes Q x and Q: x. Q x = y x. P y shows Least P Least Q proof obtain f :: a a where a. Q a f a a P (f a) using Q by moura moreover have Q (Least Q) using Q x by(rule LeastI ) ultimately show?thesis by (metis (full-types) le-cases le-less less-le-trans not-less-least) 1.1 Transfer rules context begin interpretation lifting-syntax. lemma monotone-parametric [transfer-rule]: assumes [transfer-rule]: bi-total A shows ((A ===> A ===> op =) ===> (B ===> B ===> op =) ===> (A ===> B) ===> op =) monotone monotone unfolding monotone-def [abs-def ] by transfer-prover lemma fun-ord-parametric [transfer-rule]: assumes [transfer-rule]: bi-total C shows ((A ===> B ===> op =) ===> (C ===> A) ===> (C ===> B) ===> op =) fun-ord fun-ord unfolding fun-ord-def [abs-def ] by transfer-prover lemma fun-lub-parametric [transfer-rule]: assumes [transfer-rule]: bi-total A bi-unique A shows ((rel-set A ===> B) ===> rel-set (C ===> A) ===> C ===> B) fun-lub fun-lub unfolding fun-lub-def [abs-def ] by transfer-prover lemma Ex1-parametric [transfer-rule]: assumes [transfer-rule]: bi-unique A bi-total A shows ((A ===> op =) ===> op =) Ex1 Ex1 unfolding Ex1-def [abs-def ] by transfer-prover end 1.2 Algebraic stuff lemma abs-diff-triangle-ineq2 : a b :: - :: ordered-ab-group-add-abs a c + c b by(rule order-trans[of - abs-diff-triangle-ineq]) simp lemma (in ordered-ab-semigroup-add) add-left-mono-trans: [ x a + b; b c ] = x a + c 8

9 by(erule order-trans)(rule add-left-mono) lemma of-nat-le-one-cancel-iff [simp]: fixes n :: nat shows real n 1 n 1 by linarith lemma (in linordered-semidom) mult-right-le: c 1 = 0 a = c a a by(subst mult.commute)(rule mult-left-le) 1.3 Option declare is-none-bind [simp] lemma None-in-map-option-image [simp]: None map-option f A None A by auto lemma Some-in-map-option-image [simp]: Some x map-option f A ( y. x = f y Some y A) by(auto intro: rev-image-eqi dest: sym) lemma case-option-collapse: case-option x (λ-. x) y = x by(simp split: option.split) lemma case-option-id: case-option None Some = id by(rule ext)(simp split: option.split) Orders on option inductive ord-option :: ( a b bool) a option b option bool for ord :: a b bool where None: ord-option ord None x Some: ord x y = ord-option ord (Some x) (Some y) inductive-simps ord-option-simps [simp]: ord-option ord None x ord-option ord x None ord-option ord (Some x) (Some y) ord-option ord (Some x) None inductive-simps ord-option-eq-simps [simp]: ord-option op = None y ord-option op = (Some x) y lemma ord-option-refli : ( y. y set-option x = ord y y) = ord-option ord x x by(cases x) simp-all lemma reflp-ord-option: reflp ord = reflp (ord-option ord) 9

10 by(simp add: reflp-def ord-option-refli ) lemma ord-option-trans: [ ord-option ord x y; ord-option ord y z; a b c. [ a set-option x; b set-option y; c set-option z; ord a b; ord b c ] = ord a c ] = ord-option ord x z by(auto elim!: ord-option.cases) lemma transp-ord-option: transp ord = transp (ord-option ord) unfolding transp-def by(blast intro: ord-option-trans) lemma antisymp-ord-option: antisymp ord = antisymp (ord-option ord) by(auto intro!: antisymi elim!: ord-option.cases dest: antisymd) lemma ord-option-chaind: Complete-Partial-Order.chain (ord-option ord) Y = Complete-Partial-Order.chain ord {x. Some x Y } by(rule chaini )(auto dest: chaind) definition lub-option :: ( a set b) a option set b option where lub-option lub Y = (if Y {None} then None else Some (lub {x. Some x Y })) lemma map-lub-option: map-option f (lub-option lub Y ) = lub-option (f lub) Y by(simp add: lub-option-def ) lemma lub-option-upper: assumes Complete-Partial-Order.chain (ord-option ord) Y x Y and lub-upper: Y x. [ Complete-Partial-Order.chain ord Y ; x Y ] = ord x (lub Y ) shows ord-option ord x (lub-option lub Y ) using assms(1 2 ) by(cases x)(auto simp add: lub-option-def intro: lub-upper[of ord-option-chaind]) lemma lub-option-least: assumes Y : Complete-Partial-Order.chain (ord-option ord) Y and upper: x. x Y = ord-option ord x y assumes lub-least: Y y. [ Complete-Partial-Order.chain ord Y ; x. x Y = ord x y ] = ord (lub Y ) y shows ord-option ord (lub-option lub Y ) y using Y by(cases y)(auto 4 3 simp add: lub-option-def intro: lub-least[of ord-option-chaind] dest: upper) lemma lub-map-option: lub-option lub (map-option f Y ) = lub-option (lub op f ) Y apply(auto simp add: lub-option-def ) apply(erule note) 10

11 apply(rule arg-cong[where f =lub]) apply(auto intro: rev-image-eqi dest: sym) done lemma ord-option-mono: [ ord-option A x y; x y. A x y = B x y ] = ord-option B x y by(auto elim: ord-option.cases) lemma ord-option-mono [mono]: ( x y. A x y B x y) = ord-option A x y ord-option B x y by(blast intro: ord-option-mono) lemma ord-option-compp: ord-option (A OO B) = ord-option A OO ord-option B by(auto simp add: fun-eq-iff elim!: ord-option.cases intro: ord-option.intros) lemma ord-option-inf : inf (ord-option A) (ord-option B) = ord-option (inf A B) (is?lhs =?rhs) proof(rule antisym) show?lhs?rhs by(auto elim!: ord-option.cases) (auto elim: ord-option-mono) lemma ord-option-map2 : ord-option ord x (map-option f y) = ord-option (λx y. ord x (f y)) x y by(auto elim: ord-option.cases) lemma ord-option-map1 : ord-option ord (map-option f x) y = ord-option (λx y. ord (f x) y) x y by(auto elim: ord-option.cases) lemma option-ord-some1-iff : option-ord (Some x) y y = Some x by(auto simp add: flat-ord-def ) We need a polymorphic constant for the empty map such that transfer-prover can use a custom transfer rule for Map.empty definition Map-empty where [simp]: Map-empty Map.empty filter for option fun filter-option :: ( a bool) a option a option where filter-option P None = None filter-option P (Some x) = (if P x then Some x else None) lemma set-filter-option [simp]: set-option (filter-option P x) = {y set-option x. P y} by(cases x) auto lemma filter-map-option: filter-option P (map-option f x) = map-option f (filter-option (P f ) x) 11

12 by(cases x) simp-all lemma is-none-filter-option [simp]: Option.is-none (filter-option P x) Option.is-none x P (the x) by(cases x) simp-all lemma filter-option-eq-some-iff [simp]: filter-option P x = Some y x = Some y P y by(cases x) auto lemma Some-eq-filter-option-iff [simp]: Some y = filter-option P x x = Some y P y by(cases x) auto lemma filter-conv-bind-option: filter-option P x = Option.bind x (λy. if P y then Some y else None) by(cases x) simp-all map for the combination of set and option fun map-option-set :: ( a b option set) a option b option set where map-option-set f None = {None} map-option-set f (Some x) = f x lemma None-in-map-option-set: None map-option-set f x None Set.bind (set-option x) f x = None by(cases x) simp-all lemma None-in-map-option-set-None [intro!]: None map-option-set f None by simp lemma None-in-map-option-set-Some [intro!]: None f x = None map-option-set f (Some x) by simp lemma Some-in-map-option-set [intro!]: Some y f x = Some y map-option-set f (Some x) by simp lemma map-option-set-singleton [simp]: map-option-set (λx. {f x}) y = {Option.bind y f } by(cases y) simp-all lemma Some-eq-bind-conv: Some y = Option.bind x f ( z. x = Some z f z = Some y) by(cases x) auto lemma map-option-set-bind: map-option-set f (Option.bind x g) = map-option-set 12

13 (map-option-set f g) x by(cases x) simp-all lemma Some-in-map-option-set-conv: Some y map-option-set f x ( z. x = Some z Some y f z) by(cases x) auto join on options definition join-option :: a option option a option where join-option x = (case x of Some y y None None) simps-of-case join-simps [simp, code]: join-option-def lemma set-join-option [simp]: set-option (join-option x) = (set-option set-option x) by(cases x)(simp-all) lemma in-set-join-option: x set-option (join-option (Some (Some x))) by simp lemma map-join-option: map-option f (join-option x) = join-option (map-option (map-option f ) x) by(cases x) simp-all lemma bind-conv-join-option: Option.bind x f = join-option (map-option f x) by(cases x) simp-all lemma join-conv-bind-option: join-option x = Option.bind x id by(cases x) simp-all context begin interpretation lifting-syntax. lemma join-option-parametric [transfer-rule]: (rel-option (rel-option R) ===> rel-option R) join-option join-option unfolding join-conv-bind-option[abs-def ] by transfer-prover end lemma join-option-eq-some [simp]: join-option x = Some y x = Some (Some y) by(cases x) simp-all lemma Some-eq-join-option [simp]: Some y = join-option x x = Some (Some y) by(cases x) auto lemma join-option-eq-none: join-option x = None x = None x = Some None by(cases x) simp-all 13

14 lemma None-eq-join-option: None = join-option x x = None x = Some None by(cases x) auto Zip on options function zip-option :: a option b option ( a b) option where zip-option (Some x) (Some y) = Some (x, y) zip-option - None = None zip-option None - = None by pat-completeness auto termination by lexicographic-order lemma zip-option-eq-some-iff [iff ]: zip-option x y = Some (a, b) x = Some a y = Some b by(cases (x, y) rule: zip-option.cases) simp-all lemma set-zip-option [simp]: set-option (zip-option x y) = set-option x set-option y by auto lemma zip-map-option1 : zip-option (map-option f x) y = map-option (apfst f ) (zip-option x y) by(cases (x, y) rule: zip-option.cases) simp-all lemma zip-map-option2 : zip-option x (map-option g y) = map-option (apsnd g) (zip-option x y) by(cases (x, y) rule: zip-option.cases) simp-all lemma map-zip-option: map-option (map-prod f g) (zip-option x y) = zip-option (map-option f x) (map-option g y) by(simp add: zip-map-option1 zip-map-option2 option.map-comp apfst-def apsnd-def o-def prod.map-comp) lemma zip-conv-bind-option: zip-option x y = Option.bind x (λx. Option.bind y (λy. Some (x, y))) by(cases (x, y) rule: zip-option.cases) simp-all context begin interpretation lifting-syntax. lemma zip-option-parametric [transfer-rule]: (rel-option R ===> rel-option Q ===> rel-option (rel-prod R Q)) zip-option zip-option unfolding zip-conv-bind-option[abs-def ] by transfer-prover end 14

15 1.4 Predicator for function relator with relation op = on the domain definition fun-eq-pred :: ( b bool) ( a b) bool where fun-eq-pred P f ( x. P (f x)) context begin interpretation lifting-syntax. lemma fun-rel-eq-oo: (op = ===> A) OO (op = ===> B) = (op = ===> A OO B) by(clarsimp simp add: rel-fun-def fun-eq-iff relcompp.simps) metis lemma Domainp-fun-rel-eq: Domainp (op = ===> A) = fun-eq-pred (Domainp A) by(simp add: rel-fun-def Domainp.simps fun-eq-pred-def fun-eq-iff ) metis lemma Domainp-eq: Domainp op = = (λ-. True) by(simp add: Domainp.simps fun-eq-iff ) lemma invariant-true: eq-onp (λ-. True) = op = by(simp add: eq-onp-def ) lemma fun-eq-invariant [relator-eq-onp]: op = ===> eq-onp P = eq-onp (fun-eq-pred P) by(simp add: rel-fun-def eq-onp-def fun-eq-iff fun-eq-pred-def ) metis lemma Quotient-set-rel-eq: assumes Quotient R Abs Rep T shows (rel-set T ===> rel-set T ===> op =) (rel-set R) op = proof(rule rel-funi iffi )+ fix A B C D assume AB: rel-set T A B and CD: rel-set T C D have : x y. R x y = (T x (Abs x) T y (Abs y) Abs x = Abs y) a b. T a b = Abs a = b using assms unfolding Quotient-alt-def by simp-all { assume [simp]: B = D thus rel-set R A C by(auto 4 4 intro!: rel-seti dest: rel-setd1 [OF AB, simplified] rel-setd2 [OF AB, simplified] rel-setd2 [OF CD] rel-setd1 [OF CD] simp add: elim!: rev-bexi ) next assume AC : rel-set R A C show B = D apply safe apply(drule rel-setd2 [OF AB], erule bexe) apply(drule rel-setd1 [OF AC ], erule bexe) apply(drule rel-setd1 [OF CD], erule bexe) apply(simp add: ) apply(drule rel-setd2 [OF CD], erule bexe) apply(drule rel-setd2 [OF AC ], erule bexe) 15

16 apply(drule rel-setd1 [OF AB], erule bexe) apply(simp add: ) done } lemma fun-eq-pred-parametric [transfer-rule]: assumes [transfer-rule]: bi-total A shows ((B ===> op =) ===> (A ===> B) ===> op =) fun-eq-pred fun-eq-pred unfolding fun-eq-pred-def [abs-def ] by transfer-prover end 1.5 A relator for sets that treats sets like predicates context begin interpretation lifting-syntax. definition rel-pred :: ( a b bool) a set b set bool where rel-pred R A B = (R ===> op =) (λx. x A) (λy. y B) lemma rel-predi : (R ===> op =) (λx. x A) (λy. y B) = rel-pred R A B by(simp add: rel-pred-def ) lemma rel-predd: [ rel-pred R A B; R x y ] = x A y B by(simp add: rel-pred-def rel-fun-def ) lemma Collect-parametric [transfer-rule]: ((A ===> op =) ===> rel-pred A) Collect Collect by(simp add: rel-funi rel-predi ) end 1.6 List of a given length inductive-set nlists :: a set nat a list set for A n where nlists: [ set xs A; length xs = n ] = xs nlists A n hide-fact (open) nlists lemma nlists-alt-def : nlists A n = {xs. set xs A length xs = n} by(auto simp add: nlists.simps) lemma nlists-empty: nlists {} n = (if n = 0 then {[]} else {}) by(auto simp add: nlists-alt-def ) lemma nlists-empty-gt0 [simp]: n > 0 = nlists {} n = {} by(simp add: nlists-empty) lemma nlists-0 [simp]: nlists A 0 = {[]} by(auto simp add: nlists-alt-def ) 16

17 lemma Cons-in-nlists-Suc [simp]: x # xs nlists A (Suc n) x A xs nlists A n by(simp add: nlists-alt-def ) lemma Nil-in-nlists [simp]: [] nlists A n n = 0 by(auto simp add: nlists-alt-def ) lemma Cons-in-nlists-iff : x # xs nlists A n ( n. n = Suc n x A xs nlists A n ) by(cases n) simp-all lemma in-nlists-suc-iff : xs nlists A (Suc n) ( x xs. xs = x # xs x A xs nlists A n) by(cases xs) simp-all lemma nlists-suc: nlists A (Suc n) = ( x A. op # x nlists A n) by(auto 4 3 simp add: in-nlists-suc-iff intro: rev-image-eqi ) lemma replicate-in-nlists [simp, intro]: x A = replicate n x nlists A n by(simp add: nlists-alt-def set-replicate-conv-if ) lemma nlists-eq-empty-iff [simp]: nlists A n = {} n > 0 A = {} using replicate-in-nlists by(cases n)(auto) lemma finite-nlists [simp]: finite A = finite (nlists A n) by(induction n)(simp-all add: nlists-suc) lemma finite-nlistsd: assumes finite (nlists A n) shows finite A n = 0 proof(rule disjci ) assume n 0 then obtain n where n: n = Suc n by(cases n)auto then have A = hd nlists A n by(auto 4 4 simp add: nlists-suc intro: rev-image-eqi rev-bexi ) also have finite... using assms.. finally show finite A. lemma finite-nlists-iff : finite (nlists A n) finite A n = 0 by(auto dest: finite-nlistsd) lemma card-nlists: card (nlists A n) = card A ˆ n proof(induction n) case (Suc n) have card ( x A. op # x nlists A n) = card A card (nlists A n) proof(cases finite A) case True then show?thesis by(subst card-un-disjoint)(auto simp add: card-image 17

18 inj-on-def ) next case False hence finite ( x A. op # x nlists A n) unfolding nlists-suc[symmetric] by(auto dest: finite-nlistsd) then show?thesis using False by simp then show?case using Suc.IH by(simp add: nlists-suc) simp lemma in-nlists-univ : xs nlists UNIV n length xs = n by(simp add: nlists-alt-def ) 1.7 Corecursor for products definition corec-prod :: ( s a) ( s b) s a b where corec-prod f g = (λs. (f s, g s)) lemma corec-prod-apply: corec-prod f g s = (f s, g s) by(simp add: corec-prod-def ) lemma corec-prod-sel [simp]: shows fst-corec-prod: fst (corec-prod f g s) = f s and snd-corec-prod: snd (corec-prod f g s) = g s by(simp-all add: corec-prod-apply) lemma apfst-corec-prod [simp]: apfst h (corec-prod f g s) = corec-prod (h f ) g s by(simp add: corec-prod-apply) lemma apsnd-corec-prod [simp]: apsnd h (corec-prod f g s) = corec-prod f (h g) s by(simp add: corec-prod-apply) lemma map-corec-prod [simp]: map-prod f g (corec-prod h k s) = corec-prod (f h) (g k) s by(simp add: corec-prod-apply) lemma split-corec-prod [simp]: case-prod h (corec-prod f g s) = h (f s) (g s) by(simp add: corec-prod-apply) 1.8 If named-theorems if-distribs Distributivity theorems for If lemma if-mono-cong: [b = x x ; b = y y ] = If b x y If b x y by simp end theory PF-Set imports Main begin 18

19 lemma (in complete-lattice) lattice-partial-function-definition: partial-function-definitions (op ) Sup by(unfold-locales)(auto intro: Sup-upper Sup-least) interpretation set: partial-function-definitions op Union by(rule lattice-partial-function-definition) lemma fun-lub-sup: fun-lub Sup = (Sup :: - - :: complete-lattice) by(fastforce simp add: fun-lub-def fun-eq-iff Sup-fun-def intro: Sup-eqI SUP-upper SUP-least) lemma set-admissible: set.admissible (λf :: a b set. x y. y f x P x y) by(rule ccpo.admissiblei )(auto simp add: fun-lub-sup) abbreviation mono-set monotone (fun-ord op ) op lemma fixp-induct-set-scott: fixes F :: c c and U :: c b a set and C :: ( b a set) c and P :: b a bool and x and y assumes mono: x. mono-set (λf. U (F (C f )) x) and eq: f C (ccpo.fixp (fun-lub Sup) (fun-ord op ) (λf. U (F (C f )))) and inverse2 : f. U (C f ) = f and step: f x y. [ x y. y U f x = P x y; y U (F f ) x ] = P x y and enforce-variable-ordering: x = x and elem: y U f x shows P x y using step elem set.fixp-induct-uc[of U F C, OF mono eq inverse2 set-admissible, of P] by blast lemma fixp-sup-le: defines le (op :: - :: complete-lattice -) shows ccpo.fixp Sup le = ccpo-class.fixp proof have class.ccpo Sup le op < unfolding le-def by unfold-locales thus?thesis by(simp add: ccpo.fixp-def fixp-def ccpo.iterates-def iterates-def ccpo.iteratesp-def iteratesp-def fun-eq-iff le-def ) lemma fun-ord-le: fun-ord op = op by(auto simp add: fun-ord-def fun-eq-iff le-fun-def ) lemma monotone-le-le: monotone op op = mono by(simp add: monotone-def [abs-def ] mono-def [abs-def ]) 19

20 lemma fixp-induct-set: fixes F :: c c and U :: c b a set and C :: ( b a set) c and P :: b a bool and x and y assumes mono: x. mono-set (λf. U (F (C f )) x) and eq: f C (ccpo.fixp (fun-lub Sup) (fun-ord op ) (λf. U (F (C f )))) and inverse2 : f. U (C f ) = f and step: f x y. [ x. U f x = U f x; y U (F (C (inf (U f ) (λx. {y. P x y})))) x ] = P x y partial function requires a quantifier over f, so let s have a fake one and elem: y U f x shows P x y proof from mono have mono : mono (λf. U (F (C f ))) by(simp add: fun-ord-le monotone-le-le mono-def le-fun-def ) hence eq : f C (lfp (λf. U (F (C f )))) using eq unfolding fun-ord-le fun-lub-sup fixp-sup-le by(simp add: lfp-eq-fixp) let?f = C (lfp (λf. U (F (C f )))) have step : x y. [ y U (F (C (inf (U?f ) (λx. {y. P x y})))) x ] = P x y unfolding eq [symmetric] by(rule step[of refl]) let?p = λx. {y. P x y} from mono have lfp (λf. U (F (C f )))?P by(rule lfp-induct)(auto intro!: le-funi step simp add: inverse2 ) with elem show?thesis by(subst (asm) eq )(auto simp add: inverse2 le-fun-def ) declaration Partial-Function.init set.fixp-induct-uc} fixp-induct-set}) lemma [partial-function-mono]: shows insert-mono: mono-set A = mono-set (λf. insert x (A f )) and UNION-mono: [mono-set B; y. mono-set (λf. C y f )] = mono-set (λf. y B f. C y f ) and set-bind-mono: [mono-set B; y. mono-set (λf. C y f )] = mono-set (λf. Set.bind (B f ) (λy. C y f )) and Un-mono: [ mono-set A; mono-set B ] = mono-set (λf. A f B f ) and Int-mono: [ mono-set A; mono-set B ] = mono-set (λf. A f B f ) and Diff-mono1 : mono-set A = mono-set (λf. A f X ) and image-mono: mono-set A = mono-set (λf. g A f ) and vimage-mono: mono-set A = mono-set (λf. g A f ) 20

21 unfolding bind-union by(blast intro!: monotonei dest: monotoned)+ partial-function (set) test :: a list nat bool int set where test xs i j = insert 4 (test [] 0 j test [] 1 True test [] 2 False {5 } uminus test [undefined] 0 True uminus test [] 1 False) interpretation coset: partial-function-definitions op Inter by(rule complete-lattice.lattice-partial-function-definition[of dual-complete-lattice]) lemma fun-lub-inf : fun-lub Inf = (Inf :: - - :: complete-lattice) by(auto simp add: fun-lub-def fun-eq-iff Inf-fun-def intro: Inf-eqI INF-lower INF-greatest) lemma fun-ord-ge: fun-ord op = op by(auto simp add: fun-ord-def fun-eq-iff le-fun-def ) lemma coset-admissible: coset.admissible (λf :: a b set. x y. P x y y f x) by(rule ccpo.admissiblei )(auto simp add: fun-lub-inf ) abbreviation mono-coset monotone (fun-ord op ) op lemma gfp-eq-fixp: fixes f :: a :: complete-lattice a assumes f : monotone op op f shows gfp f = ccpo.fixp Inf op f proof (rule antisym) from f have f : mono f by(simp add: mono-def monotone-def ) interpret ccpo Inf op mk-less op :: a - by(rule ccpo)(rule complete-lattice.lattice-partial-function-definition[of dual-complete-lattice]) show ccpo.fixp Inf op f gfp f by(rule gfp-upperbound)(subst fixp-unfold[of f ], rule order-refl) show gfp f ccpo.fixp Inf op f by(rule fixp-lowerbound[of f ])(subst gfp-unfold[of f ], rule order-refl) lemma fixp-coinduct-set: fixes F :: c c and U :: c b a set and C :: ( b a set) c and P :: b a bool and x and y assumes mono: x. mono-coset (λf. U (F (C f )) x) and eq: f C (ccpo.fixp (fun-lub Inter) (fun-ord op ) (λf. U (F (C f )))) and inverse2 : f. U (C f ) = f 21

22 and step: f x y. [ x. U f x = U f x; P x y ] = y U (F (C (sup (λx. {y. P x y}) (U f )))) x partial function requires a quantifier over f, so let s have a fake one and elem: y / U f x shows P x y using elem proof(rule contrapos-np) have mono : monotone (op ) (op ) (λf. U (F (C f ))) and mono : mono (λf. U (F (C f ))) using mono by(simp-all add: monotone-def fun-ord-def le-fun-def mono-def ) hence eq : U f = gfp (λf. U (F (C f ))) apply(subst eq) by(simp add: fun-lub-inf fun-ord-ge gfp-eq-fixp inverse2 ) let?p = λx. {y. P x y} have?p gfp (λf. U (F (C f ))) using mono by(rule coinduct)(auto intro!: le-funi dest: step[of refl] simp add: eq ) moreover assume P x y ultimately show y U f x by(auto simp add: le-fun-def eq ) declaration Partial-Function.init coset.fixp-induct-uc} fixp-coinduct-set}) abbreviation mono-set monotone (fun-ord op ) op lemma [partial-function-mono]: shows insert-mono : mono-set A = mono-set (λf. insert x (A f )) and UNION-mono : [mono-set B; y. mono-set (λf. C y f )] = mono-set (λf. y B f. C y f ) and set-bind-mono : [mono-set B; y. mono-set (λf. C y f )] = mono-set (λf. Set.bind (B f ) (λy. C y f )) and Un-mono : [ mono-set A; mono-set B ] = mono-set (λf. A f B f ) and Int-mono : [ mono-set A; mono-set B ] = mono-set (λf. A f B f ) unfolding bind-union by(blast intro!: monotonei dest: monotoned)+ partial-function (coset) test2 :: nat nat set where test2 x = insert x (test2 (Suc x)) lemma test2-coinduct: assumes P x y and : x y. P x y = y = x (P (Suc x) y y test2 (Suc x)) shows y test2 x using P x y apply(rule contrapos-pp) apply(erule test2.raw-induct[rotated]) 22

23 apply(simp add: ) done end theory Monad-Locale imports Misc /src/hol/library/monad-syntax begin context begin interpretation lifting-syntax. lemma rel-fun-conversep: (Aˆ 1 ===> Bˆ 1 ) = (A ===> B)ˆ 1 by(auto simp add: rel-fun-def fun-eq-iff ) end lemma left-unique-grp [iff ]: left-unique (BNF-Def.Grp A f ) inj-on f A unfolding Grp-def left-unique-def by(auto simp add: inj-on-def ) lemma right-unique-grp [simp, intro!]: right-unique (BNF-Def.Grp A f ) by(simp add: Grp-def right-unique-def ) lemma bi-unique-grp [iff ]: bi-unique (BNF-Def.Grp A f ) inj-on f A by(simp add: bi-unique-alt-def ) lemma left-total-grp [iff ]: left-total (BNF-Def.Grp A f ) A = UNIV by(auto simp add: left-total-def Grp-def ) lemma right-total-grp [iff ]: right-total (BNF-Def.Grp A f ) f A = UNIV by(auto simp add: right-total-def BNF-Def.Grp-def image-def ) lemma bi-total-grp [iff ]: bi-total (BNF-Def.Grp A f ) A = UNIV surj f by(auto simp add: bi-total-alt-def ) lemma left-unique-vimage2p [simp]: [ left-unique P; inj f ] = left-unique (BNF-Def.vimage2p f g P) unfolding vimage2p-grp by(intro left-unique-oo) simp-all lemma right-unique-vimage2p [simp]: [ right-unique P; inj g ] = right-unique (BNF-Def.vimage2p f g P) unfolding vimage2p-grp by(intro right-unique-oo) simp-all lemma bi-unique-vimage2p [simp]: [ bi-unique P; inj f ; inj g ] = bi-unique (BNF-Def.vimage2p f g P) unfolding bi-unique-alt-def by simp 23

24 lemma left-total-vimage2p [simp]: [ left-total P; surj g ] = left-total (BNF-Def.vimage2p f g P) unfolding vimage2p-grp by(intro left-total-oo) simp-all lemma right-total-vimage2p [simp]: [ right-total P; surj f ] = right-total (BNF-Def.vimage2p f g P) unfolding vimage2p-grp by(intro right-total-oo) simp-all lemma bi-total-vimage2p [simp]: [ bi-total P; surj f ; surj g ] = bi-total (BNF-Def.vimage2p f g P) unfolding bi-total-alt-def by simp lemma vimage2p-eq [simp]: inj f = BNF-Def.vimage2p f f op = = op = by(auto simp add: vimage2p-def fun-eq-iff inj-on-def ) lemma vimage2p-conversep: BNF-Def.vimage2p f g Rˆ 1 = (BNF-Def.vimage2p g f R)ˆ 1 by(simp add: vimage2p-def fun-eq-iff ) named-theorems record-parametricity unfold theorems for record-parametricity named-theorems record-typedef type definition theorems for record parametricity context begin interpretation lifting-syntax. lemma vimage2p-abs-rep: assumes h: (Y ===> X1 ===> X2 ) h1 h2 and 1 : type-definition f1 g1 UNIV and 2 : type-definition f2 g2 UNIV shows (Y ===> BNF-Def.vimage2p f1 f2 X1 ===> BNF-Def.vimage2p f1 f2 X2 ) (λx. g1 (h1 x f1 )) (λx. g2 (h2 x f2 )) apply(rule rel-funi ) apply(rule comp-transfer[then rel-fund, THEN rel-fund]) apply(rule Abs-transfer[OF 1 2 ]) apply(rule comp-transfer[then rel-fund, THEN rel-fund]) apply(erule rel-fund[of h]) apply(rule vimage2p-rel-fun) done lemma apsnd-parametric [transfer-rule]: ((A ===> B) ===> rel-prod C A ===> rel-prod C B) apsnd apsnd unfolding apsnd-def [abs-def ] by transfer-prover end method record-parametricity uses unfold typedefs = (unfold unfold record-parametricity iso-tuple-fst-def iso-tuple-snd-def tuple-iso-tuple-def iso-tuple-fst-update-def [abs-def ] iso-tuple-snd-update-def [abs-def ] 24

25 iso-tuple-cons-def repr.simps abst.simps o-assoc[symmetric] id-o o-id, ((assumption rule typedefs record-typedef fst-transfer[then comp-transfer[then rel-fund, THEN rel-fund]] snd-transfer[then comp-transfer[then rel-fund, THEN rel-fund]] comp-transfer[then rel-fund, THEN rel-fund] apfst-parametric apsnd-parametric curry-transfer[then rel-fund, THEN rel-fund] curry-transfer[then rel-fund, THEN rel-fund, THEN rel-fund, OF id-transfer] vimage2p-abs-rep Abs-transfer vimage2p-rel-fun rel-funi )+)? ) attribute-setup locale-witness = Scan.succeed Locale.witness-add 2 Monomorphic monads and monad transformers 2.1 Plain monads record ( a, m) monad = bind :: m ( a m) m return :: a m context begin local-setup Local-Theory.map-background-naming (Name-Space.mandatory-path monad) definition lift :: ( a, m, more) monad-ext ( a a) m m where lift m f x = bind m x (return m f ) end locale monad-syntax = fixes m :: ( a, m, more) monad-scheme begin adhoc-overloading Monad-Syntax.bind bind m end locale monad = fixes m :: ( a, m, more) monad-scheme assumes bind-assoc: (x :: m) f g. bind m (bind m x f ) g = bind m x (λy. bind m (f y) g) and return-bind: x f. bind m (return m x) f = f x and bind-return: x. bind m x (return m) = x lemma inj-rep-monad-ext [simp]: inj Rep-monad-ext by(simp add: inj-on-def Rep-monad-ext-inject) lemma surj-rep-monad-ext [simp]: surj Rep-monad-ext by(rule surji [where f =Abs-monad-ext])(simp add: Abs-monad-ext-inverse) 25

26 context begin interpretation lifting-syntax. definition rel-monad-ext :: ( a b bool) ( m1 m2 bool) ( more1 more2 bool) ( a, m1, more1 ) monad-ext ( b, m2, more2 ) monad-ext bool where [record-parametricity]: rel-monad-ext A M X = BNF-Def.vimage2p Rep-monad-ext Rep-monad-ext (rel-prod (rel-prod (M ===> (A ===> M ) ===> M ) (A ===> M )) X ) declare monad-ext-tuple-iso-def [record-parametricity] monad.select-defs [record-parametricity] monad.update-defs [record-parametricity] type-definition-monad-ext [record-typedef ] abbreviation rel-monad :: ( a b bool) ( m1 m2 bool) ( a, m1 ) monad ( b, m2 ) monad bool where rel-monad A M == rel-monad-ext A M op = lemma rel-monad-ext-conversep: rel-monad-ext Aˆ 1 Mˆ 1 Xˆ 1 = (rel-monad-ext A M X )ˆ 1 unfolding rel-monad-ext-def rel-fun-conversep prod.rel-conversep vimage2p-conversep.. lemma left-unique-rel-monad-ext [transfer-rule]: [ left-unique X ; left-unique A; left-total A; left-unique M ; left-total M ] = left-unique (rel-monad-ext A M X ) unfolding rel-monad-ext-def by(simp add: prod.left-unique-rel left-unique-fun left-total-fun) lemma right-unique-rel-monad-ext [transfer-rule]: [ right-unique X ; right-unique A; right-total A; right-unique M ; right-total M ] = right-unique (rel-monad-ext A M X ) unfolding rel-monad-ext-def by(simp add: prod.right-unique-rel right-unique-fun right-total-fun) lemma left-total-rel-monad-ext [transfer-rule]: [ left-total X ; left-unique A; left-total A; left-unique M ; left-total M ] = left-total (rel-monad-ext A M X ) unfolding rel-monad-ext-def by(simp add: prod.left-total-rel left-total-fun left-unique-fun) lemma right-total-rel-monad-ext [transfer-rule]: [ right-total X ; right-unique A; right-total A; right-unique M ; right-total M ] = right-total (rel-monad-ext A M X ) unfolding rel-monad-ext-def 26

27 by(simp add: prod.right-total-rel right-total-fun right-unique-fun) lemma rel-monad-ext-eq [relator-eq]: rel-monad-ext op = op = op = = op = unfolding rel-monad-ext-def rel-fun-eq prod.rel-eq by(simp add: eq-oo eq-alt[symmetric]) lemma bind-monad-parametric [transfer-rule]: (rel-monad-ext A M X ===> M ===> (A ===> M ) ===> M ) bind bind by(record-parametricity) lemma return-monad-parametric [transfer-rule]: (rel-monad-ext A M X ===> A ===> M ) return return by(record-parametricity) lemma more-monad-parametric [transfer-rule]: (rel-monad-ext A M X ===> X ) more more by(record-parametricity) lemma lift-monad-parametric [transfer-rule]: (rel-monad-ext A M X ===> (A ===> A) ===> M ===> M ) monad.lift monad.lift unfolding monad.lift-def [abs-def ] by transfer-prover lemma bind-update-parametric [transfer-rule]: (((M ===> (A ===> M ) ===> M ) ===> (M ===> (A ===> M ) ===> M )) ===> rel-monad-ext A M X ===> rel-monad-ext A M X ) bind-update bind-update by record-parametricity lemma return-update-parametric [transfer-rule]: (((A ===> M ) ===> (A ===> M )) ===> rel-monad-ext A M X ===> rel-monad-ext A M X ) return-update return-update by record-parametricity lemma more-update-monad-parametric [transfer-rule]: ((X ===> X ) ===> rel-monad-ext A M X ===> rel-monad-ext A M X ) more-update more-update by record-parametricity lemma monad-ext-parametric [transfer-rule]: ((M ===> (A ===> M ) ===> M ) ===> (A ===> M ) ===> X ===> rel-monad-ext A M X ) monad-ext monad-ext by(record-parametricity unfold: monad-ext-def [abs-def ]) lemma make-monad-parametric [transfer-rule]: ((M ===> (A ===> M ) ===> M ) ===> (A ===> M ) ===> rel-monad 27

28 A M ) make make unfolding monad.defs[abs-def ] by transfer-prover lemma extend-monad-parametric [transfer-rule]: (rel-monad A M ===> X ===> rel-monad-ext A M X ) extend extend unfolding monad.defs[abs-def ] by transfer-prover lemma fields-monad-parametric [transfer-rule]: ((M ===> (A ===> M ) ===> M ) ===> (A ===> M ) ===> rel-monad A M ) fields fields unfolding monad.defs[abs-def ] by transfer-prover lemma truncate-monad-parametric [transfer-rule]: (rel-monad-ext A M X ===> rel-monad A M ) truncate truncate unfolding monad.defs[abs-def ] by transfer-prover end Identity monad definition identity-monad-ext :: more ( a, a, more) monad-ext where more. identity-monad-ext more = ( bind = λx f. f x, return = λx. x,... = more ) abbreviation identity-monad :: ( a, a) monad where identity-monad identity-monad-ext () lemma identity-sel [simp]: shows bind-identity: more. bind (identity-monad-ext more) = (λx f. f x) and return-identity: more. return (identity-monad-ext more) = (λx. x) and more-identity: more. monad.more (identity-monad-ext more) = more by(simp-all add: identity-monad-ext-def ) lemma monad-identity [simp, intro!, locale-witness]: more. monad (identity-monad-ext more) by unfold-locales simp-all context begin interpretation lifting-syntax. lemma identity-monad-parametric [transfer-rule]: (X ===> rel-monad-ext A A X ) identity-monad-ext identity-monad-ext unfolding identity-monad-ext-def [abs-def ] by transfer-prover end 2.2 Monads with failure record ( a, m) monad-fail = ( a, m) monad + fail :: m copy-bnf ( m, mpre) monad-fail-ext 28

29 declare type-definition-monad-fail-ext [record-typedef ] rel-monad-fail-ext-def [record-parametricity] monad-fail-ext-tuple-iso-def [record-parametricity] monad-fail.select-defs [record-parametricity] monad-fail.update-defs [record-parametricity] abbreviation rel-monad-fail-scheme :: ( a b bool) ( m1 m2 bool) ( more1 more2 bool) ( a, m1, more1 ) monad-fail-scheme ( b, m2, more2 ) monad-fail-scheme bool where rel-monad-fail-scheme A M X rel-monad-ext A M (rel-monad-fail-ext M X ) abbreviation rel-monad-fail :: ( a b bool) ( m1 m2 bool) ( a, m1 ) monad-fail ( b, m2 ) monad-fail bool where rel-monad-fail A M rel-monad-fail-scheme A M op = context begin interpretation lifting-syntax. lemma fail-monad-parametric [transfer-rule]: (rel-monad-fail-scheme A M X ===> M ) fail fail by record-parametricity lemma more-monad-fail-parametric [transfer-rule]: (rel-monad-fail-scheme A M X ===> X ) more more by record-parametricity lemma fail-update-parametric [transfer-rule]: ((M ===> M ) ===> rel-monad-fail-scheme A M X ===> rel-monad-fail-scheme A M X ) fail-update fail-update by record-parametricity lemma more-update-monad-fail-parametric [transfer-rule]: ((X ===> X ) ===> rel-monad-fail-scheme A M X ===> rel-monad-fail-scheme A M X ) more-update more-update by record-parametricity lemma monad-fail-ext-parametric [transfer-rule]: (M ===> X ===> rel-monad-fail-ext M X ) monad-fail-ext monad-fail-ext by(record-parametricity unfold: monad-fail-ext-def [abs-def ]) lemma make-monad-fail-parametric [transfer-rule]: ((M ===> (A ===> M ) ===> M ) ===> (A ===> M ) ===> M ===> rel-monad-fail A M ) make make unfolding monad-fail.defs[abs-def ] by transfer-prover lemma extend-monad-fail-parametric [transfer-rule]: 29

30 (rel-monad-fail A M ===> X ===> rel-monad-fail-scheme A M X ) extend extend unfolding monad-fail.defs[abs-def ] by transfer-prover lemma fields-monad-fail-parametric [transfer-rule]: (M ===> rel-monad-fail-ext M op =) fields fields unfolding monad-fail.defs[abs-def ] by transfer-prover lemma truncate-monad-fail-parametric [transfer-rule]: (rel-monad-fail-scheme A M X ===> rel-monad-fail A M ) truncate truncate unfolding monad-fail.defs[abs-def ] by transfer-prover end context begin local-setup Local-Theory.map-background-naming (Name-Space.mandatory-path monad) definition bind-option :: ( a, m, more) monad-fail-scheme b option ( b m) m where bind-option m x f = (case x of None fail m Some x f x ) simps-of-case bind-option-simps [simp]: bind-option-def interpretation lifting-syntax. lemma bind-option-parametric [transfer-rule]: (rel-monad-fail-scheme A M X ===> rel-option B ===> (B ===> M ) ===> M ) bind-option bind-option unfolding bind-option-def [abs-def ] by transfer-prover end locale monad-fail-syntax = monad-syntax m for m :: ( a, m, more) monad-fail-scheme begin adhoc-overloading Monad-Syntax.bind monad.bind-option m end locale monad-fail = monad m for m :: ( a, m, more) monad-fail-scheme + assumes fail-bind: f. bind m (fail m) f = fail m 2.3 The option monad definition option-monad-ext :: more ( a, a option, more) monad-fail-scheme where more. option-monad-ext more = ( bind = Option.bind, return = Some, fail = None,... = more ) abbreviation option-monad :: ( a, a option) monad-fail where option-monad option-monad-ext () lemma option-monad-sel [simp]: shows bind-option-monad: more. bind (option-monad-ext more) = Option.bind and return-option-monad: more. return (option-monad-ext more) = Some 30

31 and fail-option-monad: more. fail (option-monad-ext more) = None and more-option-monad: more. monad-fail.more (option-monad-ext more) = more by(simp-all add: option-monad-ext-def ) lemma option-monad [simp, intro!, locale-witness]: more. monad (option-monad-ext more) by unfold-locales simp-all lemma option-monad-fail [simp, intro!, locale-witness]: more. monad-fail (option-monad-ext more) by unfold-locales simp context begin interpretation lifting-syntax. lemma option-monad-ext-parametric [transfer-rule]: (X ===> rel-monad-fail-scheme A (rel-option A) X ) option-monad-ext option-monad-ext unfolding option-monad-ext-def [abs-def ] by transfer-prover end lemma bind-option-option [simp]: more. monad.bind-option (option-monad-ext more) = Option.bind by(simp add: monad.bind-option-def fun-eq-iff split: option.split) lemma monad-lift-option [simp]: more. monad.lift option-monad = map-option by(simp add: monad.lift-def map-conv-bind-option fun-eq-iff ) 2.4 The set monad definition set-monad-ext :: more ( a, a set, more) monad-fail-scheme where more. set-monad-ext more = ( bind = UNION, return = λx. {x}, fail = {},... = more ) abbreviation set-monad :: ( a, a set) monad-fail where set-monad set-monad-ext () lemma set-monad-sel [simp]: shows bind-set-monad: more. bind (set-monad-ext more) = UNION and return-set-monad: more. return (set-monad-ext more) = (λx. {x}) and fail-set-monad: more. fail (set-monad-ext more) = {} and more-set-monad: more. monad-fail.more (set-monad-ext more) = more by(simp-all add: set-monad-ext-def ) lemma set-monad [simp, intro!, locale-witness]: more. monad (set-monad-ext more) by unfold-locales auto lemma set-monad-fail [simp, intro!, locale-witness]: more. monad-fail (set-monad-ext more) by unfold-locales simp 31

32 context begin interpretation lifting-syntax. lemma set-monad-ext-parametric [transfer-rule]: (X ===> rel-monad-fail-scheme A (rel-set A) X ) set-monad-ext set-monad-ext unfolding set-monad-ext-def [abs-def ] by transfer-prover end lemma bind-option-set [simp]: more. monad.bind-option (set-monad-ext more) = (λx. UNION (set-option x)) by(simp add: monad.bind-option-def fun-eq-iff split: option.split) lemma monad-lift-set [simp]: more. monad.lift (set-monad-ext more) = image by(simp add: monad.lift-def image-eq-un o-def fun-eq-iff ) 2.5 Monad homomorphism locale monad-hom = m1 : monad m1 + m2 : monad m2 for m1 :: ( a, m1, more1 ) monad-scheme and m2 :: ( a, m2, more2 ) monad-scheme and h :: m1 m2 + assumes hom-return: x. h (return m1 x) = return m2 x and hom-bind: x f. h (bind m1 x f ) = bind m2 (h x) (h f ) begin lemma hom-lift: h (monad.lift m1 f x) = monad.lift m2 f (h x) by(simp add: monad.lift-def hom-bind hom-return o-def ) end locale monad-fail-hom = m1 : monad-fail m1 + m2 : monad-fail m2 + monad-hom m1 m2 h for m1 :: ( a, m1, more1 ) monad-fail-scheme and m2 :: ( a, m2, more2 ) monad-fail-scheme and h :: m1 m2 + assumes hom-fail: h (fail m1 ) = fail m2 begin lemma hom-bind-option: h (monad.bind-option m1 x f ) = monad.bind-option m2 x (h f ) by(cases x)(simp-all add: hom-fail) end 32

33 3 Monad transformers 3.1 The state monad transformer datatype ( s, m) ST = ST (run-state: s m) for rel: rel-st hide-const (open) ST primrec bind-state :: ( a s, m, more) monad-scheme ( s, m) ST ( a ( s, m) ST ) ( s, m) ST where bind-state m (ST.ST x) f = ST.ST (λs. bind m (x s) (λ(a, s ). run-state (f a) s )) lemma run-bind-state [simp]: run-state (bind-state m x f ) s = bind m (run-state x s) (λ(a, s ). run-state (f a) s ) by(cases x) simp definition return-state :: ( a s, m, more) monad-scheme a ( s, m) ST where return-state m x = ST.ST (λs. return m (x, s)) lemma run-return-state [simp]: run-state (return-state m x) s = return m (x, s) by(simp add: return-state-def ) context fixes m :: ( a s, m, more) monad-scheme assumes monad m begin interpretation monad m by fact lemma bind-state-assoc [simp]: bind-state m (bind-state m x f ) g = bind-state m x (λy. bind-state m (f y) g) by(rule ST.expand ext)+(simp add: bind-assoc split-def ) lemma return-bind-state [simp]: bind-state m (return-state m x) f = f x by(rule ST.expand ext)+(simp add: return-bind) lemma bind-return-state [simp]: bind-state m x (return-state m) = x by(rule ST.expand ext)+(simp add: bind-return) end context fixes m :: ( a s, m, more) monad-fail-scheme begin definition fail-state :: ( s, m) ST where fail-state = ST.ST (λ-. fail m) lemma run-fail-state [simp]: run-state fail-state s = fail m by(simp add: fail-state-def ) 33

Boolean Expression Checkers

Boolean Expression Checkers Boolean Expression Checkers Tobias Nipkow May 27, 2015 Abstract This entry provides executable checkers for the following properties of boolean expressions: satisfiability, tautology and equivalence. Internally,

More information

The Divergence of the Prime Harmonic Series

The Divergence of the Prime Harmonic Series The Divergence of the Prime Harmonic Series Manuel Eberl April 17, 2016 Abstract In this work, we prove the lower bound ln(h n ) ln( 5 3 ) for the partial sum of the Prime Harmonic series and, based on

More information

Imperative Insertion Sort

Imperative Insertion Sort Imperative Insertion Sort Christian Sternagel October 11, 2017 Contents 1 Looping Constructs for Imperative HOL 1 1.1 While Loops............................ 1 1.2 For Loops.............................

More information

Lattices and Orders in Isabelle/HOL

Lattices and Orders in Isabelle/HOL Lattices and Orders in Isabelle/HOL Markus Wenzel TU München October 8, 2017 Abstract We consider abstract structures of orders and lattices. Many fundamental concepts of lattice theory are developed,

More information

Stone Algebras. Walter Guttmann. September 6, 2016

Stone Algebras. Walter Guttmann. September 6, 2016 Stone Algebras Walter Guttmann September 6, 2016 Abstract A range of algebras between lattices and Boolean algebras generalise the notion of a complement. We develop a hierarchy of these pseudo-complemented

More information

Free Groups. Joachim Breitner. April 17, 2016

Free Groups. Joachim Breitner. April 17, 2016 Free Groups Joachim Breitner April 17, 2016 Abstract Free Groups are, in a sense, the most generic kind of group. They are defined over a set of generators with no additional relations in between them.

More information

Menger's Theorem. Christoph Dittmann August 16, Contents

Menger's Theorem. Christoph Dittmann August 16, Contents Menger's Theorem Christoph Dittmann isabelle@christoph-d.de August 16, 2018 We present a formalization of Menger's Theorem for directed and undirected graphs in Isabelle/HOL. This well-known result shows

More information

The Tortoise and the Hare Algorithm

The Tortoise and the Hare Algorithm The Tortoise and the Hare Algorithm Peter Gammie April 17, 2016 Abstract We formalize the Tortoise and Hare cycle-finding algorithm ascribed to Floyd by Knuth (1981, p7, exercise 6), and an improved version

More information

Gödel s Incompleteness Theorems. Lawrence C. Paulson

Gödel s Incompleteness Theorems. Lawrence C. Paulson Gödel s Incompleteness Theorems Lawrence C. Paulson October 10, 2017 Abstract Gödel s two incompleteness theorems [2] are formalised, following a careful presentation by Świerczkowski [3], in the theory

More information

Hall s Marriage Theorem

Hall s Marriage Theorem Hall s Marriage Theorem Dongchen Jiang and Tobias Nipkow February 16, 2013 Abstract A proof of Hall s Marriage Theorem due to Halmos and Vaughan [1]. theory Marriage imports Main begin theorem marriage-necessary:

More information

Positional Determinacy of Parity Games

Positional Determinacy of Parity Games Positional Determinacy of Parity Games Christoph Dittmann christoph.dittmann@tu-berlin.de April 17, 2016 We present a formalization of parity games (a two-player game on directed graphs) and a proof of

More information

Syntax and semantics of a GPU kernel programming language

Syntax and semantics of a GPU kernel programming language Syntax and semantics of a GPU kernel programming language John Wickerson April 17, 2016 Abstract This document accompanies the article The Design and Implementation of a Verification Technique for GPU

More information

Concrete Semantics. Tobias Nipkow & Gerwin Klein. October 8, 2017

Concrete Semantics. Tobias Nipkow & Gerwin Klein. October 8, 2017 Concrete Semantics Tobias Nipkow & Gerwin Klein October 8, 2017 Abstract This document presents formalizations of the semantics of a simple imperative programming language together with a number of applications:

More information

Markov chains and Markov decision processes in Isabelle/HOL

Markov chains and Markov decision processes in Isabelle/HOL Journal of Automated Reasoning manuscript No. (will be inserted by the editor) Markov chains and Markov decision processes in Isabelle/HOL Johannes Hölzl Received: / Accepted: Abstract This paper presents

More information

Landau Symbols. Manuel Eberl. November 28, 2018

Landau Symbols. Manuel Eberl. November 28, 2018 Landau Symbols Manuel Eberl November 28, 2018 Contents 1 Sorting and grouping factors 1 2 Decision procedure for real functions 4 2.1 Eventual non-negativity/non-zeroness............. 4 2.2 Rewriting Landau

More information

MP 5 Program Transition Systems and Linear Temporal Logic

MP 5 Program Transition Systems and Linear Temporal Logic MP 5 Program Transition Systems and Linear Temporal Logic CS 477 Spring 2018 Revision 1.0 Assigned April 10, 2018 Due April 17, 2018, 9:00 PM Extension extend48 hours (penalty 20% of total points possible)

More information

A Formally Verified Checker of the Safe Distance Traffic Rules for Autonomous Vehicles

A Formally Verified Checker of the Safe Distance Traffic Rules for Autonomous Vehicles A Formally Verified Checker of the Safe Distance Traffic Rules for Autonomous Vehicles Albert Rizaldi, Fabian Immler February 19, 2016 Abstract The Vienna Convention on Road Traffic defines the safe distance

More information

Much Ado about Two. By Sascha Böhme. February 22, 2013

Much Ado about Two. By Sascha Böhme. February 22, 2013 Much Ado about Two By Sascha Böhme February 22, 2013 Abstract This article is an Isabelle formalisation of a paper with the same. In a similar way as Knuth s 0-1-principle for sorting algorithms, that

More information

A General Method for the Proof of Theorems on Tail-recursive Functions

A General Method for the Proof of Theorems on Tail-recursive Functions A General Method for the Proof of Theorems on Tail-recursive Functions Pasquale Noce Security Certification Specialist at Arjo Systems - Gep S.p.A. pasquale dot noce dot lavoro at gmail dot com pasquale

More information

Examples for program extraction in Higher-Order Logic

Examples for program extraction in Higher-Order Logic Examples for program extraction in Higher-Order Logic Stefan Berghofer October 10, 2011 Contents 1 Auxiliary lemmas used in program extraction examples 1 2 Quotient and remainder 2 3 Greatest common divisor

More information

Concrete Semantics. Tobias Nipkow & Gerwin Klein. October 16, 2017

Concrete Semantics. Tobias Nipkow & Gerwin Klein. October 16, 2017 Concrete Semantics Tobias Nipkow & Gerwin Klein October 16, 2017 Abstract This document presents formalizations of the semantics of a simple imperative programming language together with a number of applications:

More information

Ribbon Proofs for Separation Logic (Isabelle Formalisation)

Ribbon Proofs for Separation Logic (Isabelle Formalisation) Ribbon Proofs for Separation Logic (Isabelle Formalisation) John Wickerson April 17, 2016 Abstract This document concerns the theory of ribbon proofs: a diagrammatic proof system, based on separation logic,

More information

Matrices, Jordan Normal Forms, and Spectral Radius Theory

Matrices, Jordan Normal Forms, and Spectral Radius Theory Matrices, Jordan Normal Forms, and Spectral Radius Theory René Thiemann and Akihisa Yamada October 10, 2017 Abstract Matrix interpretations are useful as measure functions in termination proving. In order

More information

Isabelle/FOL First-Order Logic

Isabelle/FOL First-Order Logic Isabelle/FOL First-Order Logic Larry Paulson and Markus Wenzel August 15, 2018 Contents 1 Intuitionistic first-order logic 2 1.1 Syntax and axiomatic basis................... 2 1.1.1 Equality..........................

More information

Kleene Algebra with Tests and Demonic Refinement Algebras

Kleene Algebra with Tests and Demonic Refinement Algebras Kleene Algebra with Tests and Demonic Refinement Algebras Alasdair Armstrong Victor B. F. Gomes Georg Struth August 28, 2014 Abstract We formalise Kleene algebra with tests (KAT) and demonic refinement

More information

Automated Reasoning. Lecture 9: Isar A Language for Structured Proofs

Automated Reasoning. Lecture 9: Isar A Language for Structured Proofs Automated Reasoning Lecture 9: Isar A Language for Structured Proofs Jacques Fleuriot jdf@inf.ed.ac.uk Acknowledgement: Tobias Nipkow kindly provided the slides for this lecture Apply scripts unreadable

More information

The Hahn-Banach Theorem for Real Vector Spaces

The Hahn-Banach Theorem for Real Vector Spaces The Hahn-Banach Theorem for Real Vector Spaces Gertrud Bauer http://www.in.tum.de/ bauerg/ February 12, 2013 Abstract The Hahn-Banach Theorem is one of the most fundamental results in functional analysis.

More information

This is a repository copy of Theory of Designs in Isabelle/UTP. White Rose Research Online URL for this paper:

This is a repository copy of Theory of Designs in Isabelle/UTP. White Rose Research Online URL for this paper: This is a repository copy of Theory of Designs in Isabelle/UTP. White Rose Research Online URL for this paper: http://eprints.whiterose.ac.uk/129380/ Monograph: Foster, Simon David orcid.org/0000-0002-9889-9514,

More information

Cauchy s Mean Theorem and the Cauchy-Schwarz Inequality. Benjamin Porter

Cauchy s Mean Theorem and the Cauchy-Schwarz Inequality. Benjamin Porter Cauchy s Mean Theorem and the Cauchy-Schwarz Inequality Benjamin Porter March 12, 2013 Contents 1 Cauchy s Mean Theorem 3 1.1 Abstract.............................. 3 1.2 Formal proof...........................

More information

Executable Transitive Closures of Finite Relations

Executable Transitive Closures of Finite Relations Executable Transitive Closures of Finite Relations Christian Sternagel and René Thiemann August 28, 2014 Abstract We provide a generic work-list algorithm to compute the transitive closure of finite relations

More information

The Tortoise and the Hare Algorithm

The Tortoise and the Hare Algorithm The Tortoise and the Hare Algorithm Peter Gammie October 11, 2017 Abstract We formalize the Tortoise and Hare cycle-finding algorithm ascribed to Floyd by Knuth (1981, p7, exercise 6), and an improved

More information

HOL: Well-founded and Primitive Recursion

HOL: Well-founded and Primitive Recursion Dipl.-Inf. Achim D. Brucker Dr. Burkhart Wolff Computer-supported Modeling and Reasoning http://www.infsec.ethz.ch/ education/permanent/csmr/ (rev. 16814) Submission date: HOL: Well-founded and Primitive

More information

Category Theory with Adjunctions and Limits

Category Theory with Adjunctions and Limits Category Theory with Adjunctions and Limits Eugene W. Stark Department of Computer Science Stony Brook University Stony Brook, New York 11794 USA October 10, 2017 Abstract This article attempts to develop

More information

Miscellaneous HOL Examples

Miscellaneous HOL Examples Miscellaneous HOL Examples June 8, 2008 Contents 1 Foundations of HOL 6 1.1 Pure Logic............................. 6 1.1.1 Basic logical connectives................. 6 1.1.2 Extensional equality...................

More information

A Formalisation of Lehmer s Primality Criterion

A Formalisation of Lehmer s Primality Criterion A Formalisation of Lehmer s Primality Criterion By Simon Wimmer and Lars Noschinski April 17, 2016 Abstract In 1927, Lehmer presented criterions for primality, based on the converse of Fermat s litte theorem

More information

Dijkstra s Algorithm

Dijkstra s Algorithm Dijkstra s Algorithm Benedikt Nordhoff Peter Lammich April 17, 2016 Abstract We implement and prove correct Dijkstra s algorithm for the single source shortest path problem, conceived in 1956 by E. Dijkstra.

More information

Imperative Insertion Sort

Imperative Insertion Sort Imperative Insertion Sort Christian Sternagel April 17, 2016 Contents 1 Looping Constructs for Imperative HOL 1 1.1 While Loops............................ 1 1.2 For Loops.............................

More information

SAT Solver verification

SAT Solver verification SAT Solver verification By Filip Marić April 17, 2016 Abstract This document contains formall correctness proofs of modern SAT solvers. Two different approaches are used state-transition systems shallow

More information

Verification of the Deutsch-Schorr-Waite Graph Marking Algorithm using Data Refinement

Verification of the Deutsch-Schorr-Waite Graph Marking Algorithm using Data Refinement Verification of the Deutsch-Schorr-Waite Graph Marking Algorithm using Data Refinement Viorel Preoteasa and Ralph-Johan Back February 16, 2013 Abstract The verification of the Deutsch-Schorr-Waite graph

More information

Fundamental Properties of Lambda-calculus

Fundamental Properties of Lambda-calculus Fundamental Properties of Lambda-calculus Tobias Nipkow Stefan Berghofer February 12, 2013 Contents 1 Basic definitions of Lambda-calculus 4 1.1 Lambda-terms in de Bruijn notation and substitution....

More information

Miscellaneous Isabelle/Isar examples

Miscellaneous Isabelle/Isar examples Miscellaneous Isabelle/Isar examples Makarius Wenzel With contributions by Gertrud Bauer and Tobias Nipkow October 8, 2017 Abstract Isar offers a high-level (and theory) language for Isabelle. We give

More information

Rank-Nullity Theorem in Linear Algebra

Rank-Nullity Theorem in Linear Algebra Rank-Nullity Theorem in Linear Algebra By Jose Divasón and Jesús Aransay April 17, 2016 Abstract In this contribution, we present some formalizations based on the HOL-Multivariate-Analysis session of Isabelle.

More information

Inductive Definitions and Fixed Points

Inductive Definitions and Fixed Points 6. Inductive Definitions and Fixed Points 6.0 6. Inductive Definitions and Fixed Points 6.0 Chapter 6 Overview of Chapter Inductive Definitions and Fixed Points 6. Inductive Definitions and Fixed Points

More information

An Isabelle/HOL Formalization of the Textbook Proof of Huffman s Algorithm

An Isabelle/HOL Formalization of the Textbook Proof of Huffman s Algorithm An Isabelle/HOL Formalization of the Textbook Proof of Huffman s Algorithm Jasmin Christian Blanchette Institut für Informatik, Technische Universität München, Germany blanchette@in.tum.de March 12, 2013

More information

Formally Verified Computation of Enclosures of Solutions of Ordinary Differential Equations

Formally Verified Computation of Enclosures of Solutions of Ordinary Differential Equations Formally Verified Computation of Enclosures of Solutions of Ordinary Differential Equations Fabian Immler April 17, 2016 Abstract Ordinary differential equations (ODEs) are ubiquitous when modeling continuous

More information

The Factorization Algorithm of Berlekamp and Zassenhaus

The Factorization Algorithm of Berlekamp and Zassenhaus The Factorization Algorithm of Berlekamp and Zassenhaus Jose Divasón Sebastiaan Joosten René Thiemann Akihisa Yamada November 10, 2018 Abstract We formalize the Berlekamp-Zassenhaus algorithm for factoring

More information

Intersecting Chords Theorem

Intersecting Chords Theorem Intersecting Chords Theorem Lukas Bulwahn October 11, 2017 Abstract This entry provides a geometric proof of the intersecting chords theorem. The theorem states that when two chords intersect each other

More information

Computing N-th Roots using the Babylonian Method

Computing N-th Roots using the Babylonian Method Computing N-th Roots using the Babylonian Method René Thiemann May 27, 2015 Abstract We implement the Babylonian method [1] to compute n-th roots of numbers. We provide precise algorithms for naturals,

More information

The Isabelle/HOL Algebra Library

The Isabelle/HOL Algebra Library The Isabelle/HOL Algebra Library Clemens Ballarin (Editor) With contributions by Jesús Aransay, Clemens Ballarin, Stephan Hohe, Florian Kammüller and Lawrence C Paulson December 3, 2009 Contents 1 Objects

More information

Programming with Dependent Types in Coq

Programming with Dependent Types in Coq Programming with Dependent Types in Coq Matthieu Sozeau LRI, Univ. Paris-Sud - Démons Team & INRIA Saclay - ProVal Project PPS Seminar February 26th 2009 Paris, France Coq A higher-order, polymorphic logic:

More information

Total (Co)Programming with Guarded Recursion

Total (Co)Programming with Guarded Recursion Total (Co)Programming with Guarded Recursion Andrea Vezzosi Department of Computer Science and Engineering Chalmers University of Technology, Gothenburg, Sweden Types for Proofs and Programs Annual Meeting

More information

NICTA Advanced Course. Theorem Proving Principles, Techniques, Applications

NICTA Advanced Course. Theorem Proving Principles, Techniques, Applications NICTA Advanced Course Theorem Proving Principles, Techniques, Applications λ 1 CONTENT Intro & motivation, getting started with Isabelle Foundations & Principles Lambda Calculus Higher Order Logic, natural

More information

Equational Reasoning with Applicative Functors. Institute of Information Security

Equational Reasoning with Applicative Functors. Institute of Information Security Equational Reasoning with Applicative Functors Andreas Lochbihler Joshua Schneider Institute of Information Security Equational Reasoning with Applicative Functors Andreas Lochbihler Joshua Schneider Institute

More information

Functional Data Structures

Functional Data Structures Functional Data Structures with Isabelle/HOL Tobias Nipkow Fakultät für Informatik Technische Universität München 2017-2-3 1 Part II Functional Data Structures 2 Chapter 1 Binary Trees 3 1 Binary Trees

More information

Interactive Theorem Provers

Interactive Theorem Provers Interactive Theorem Provers from the perspective of Isabelle/Isar Makarius Wenzel Univ. Paris-Sud, LRI July 2014 = Isabelle λ β Isar α 1 Introduction Notable ITP systems LISP based: ACL2 http://www.cs.utexas.edu/users/moore/acl2

More information

Automated Reasoning Lecture 17: Inductive Proof (in Isabelle)

Automated Reasoning Lecture 17: Inductive Proof (in Isabelle) Automated Reasoning Lecture 17: Inductive Proof (in Isabelle) Jacques Fleuriot jdf@inf.ed.ac.uk Recap Previously: Unification and Rewriting This time: Proof by Induction (in Isabelle) Proof by Mathematical

More information

A Formal Proof of Correctness of a Distributed Presentation Software System

A Formal Proof of Correctness of a Distributed Presentation Software System A Formal Proof of Correctness of a Distributed Presentation Software System Ievgen Ivanov, Taras Panchenko 1 Taras Shevchenko National University of Kyiv, 64/13, Volodymyrska st., Kyiv, 01601, Ukraine,

More information

Amortized Complexity Verified

Amortized Complexity Verified Amortized Complexity Verified Tobias Nipkow August 28, 2014 Abstract A framework for the analysis of the amortized complexity of (functional) data structures is formalized in Isabelle/HOL and applied to

More information

Basic Propositional Logic. Inductive Theory of the Natural Numbers. Conjunction. Equivalence. Negation and Inequivalence. Implication.

Basic Propositional Logic. Inductive Theory of the Natural Numbers. Conjunction. Equivalence. Negation and Inequivalence. Implication. McMaster University COMPSCI&SFWRENG 2DM3 Dept. of Computing and Software Theorem List 4 Dr. W. Kahl 2017-12-09 The names listed here are precisely the names used in the preloaded material you are already

More information

Kleene Algebra. Alasdair Armstrong, Victor B. F. Gomes, Georg Struth and Tjark Weber. April 17, 2016

Kleene Algebra. Alasdair Armstrong, Victor B. F. Gomes, Georg Struth and Tjark Weber. April 17, 2016 Kleene Algebra Alasdair Armstrong, Victor B. F. Gomes, Georg Struth and Tjark Weber April 17, 2016 Abstract Variants of Dioids and Kleene algebras are formalised together with their most important models

More information

Program Construction and Verification Components Based on Kleene Algebra

Program Construction and Verification Components Based on Kleene Algebra Program Construction and Verification Components Based on Kleene Algebra Victor B. F. Gomes and Georg Struth October 11, 2017 Abstract Variants of Kleene algebra support program construction and verification

More information

Depending on equations

Depending on equations Depending on equations A proof-relevant framework for unification in dependent type theory Jesper Cockx DistriNet KU Leuven 3 September 2017 Unification for dependent types Unification is used for many

More information

Outline. A recursive function follows the structure of inductively-defined data.

Outline. A recursive function follows the structure of inductively-defined data. Outline A recursive function follows the structure of inductively-defined data. With lists as our example, we shall study 1. inductive definitions (to specify data) 2. recursive functions (to process data)

More information

Structuring the verification of heap-manipulating programs

Structuring the verification of heap-manipulating programs Structuring the verification of heap-manipulating programs Aleksandar Nanevski (IMDEA Madrid) Viktor Vafeiadis (MSR / Univ. of Cambridge) Josh Berdine (MSR Cambridge) Hoare/Separation Logic Hoare logic

More information

Subresultants. Sebastiaan Joosten, René Thiemann and Akihisa Yamada. October 10, 2017

Subresultants. Sebastiaan Joosten, René Thiemann and Akihisa Yamada. October 10, 2017 Subresultants Sebastiaan Joosten, René Thiemann and Akihisa Yamada October 10, 2017 Abstract We formalize the theory of subresultants and the subresultant polynomial remainder sequence as described by

More information

Code Generation for a Simple First-Order Prover

Code Generation for a Simple First-Order Prover Code Generation for a Simple First-Order Prover Jørgen Villadsen, Anders Schlichtkrull, and Andreas Halkjær From DTU Compute, Technical University of Denmark, 2800 Kongens Lyngby, Denmark Abstract. We

More information

Polynomial Factorization

Polynomial Factorization Polynomial Factorization René Thiemann and Akihisa Yamada October 11, 2017 Abstract Based on existing libraries for polynomial interpolation and matrices, we formalized several factorization algorithms

More information

Univalence for Free. Matthieu Sozeau Joint work with Nicolas Tabareau - INRIA. GT Types & Réalisabilité March 13th 2013

Univalence for Free. Matthieu Sozeau Joint work with Nicolas Tabareau - INRIA. GT Types & Réalisabilité March 13th 2013 Univalence for Free Matthieu Sozeau Joint work with Nicolas Tabareau - INRIA Project Team πr 2 INRIA Rocquencourt & PPS, Paris 7 University GT Types & Réalisabilité March 13th 2013 Univalence for Free

More information

A New Look At Generalized Rewriting in Type Theory

A New Look At Generalized Rewriting in Type Theory A New Look At Generalized Rewriting in Type Theory Matthieu Sozeau Harvard University TYPES 09 May 13th 2009 Aussois, France Generalized Rewriting Equational reasoning x = y - x + 1 ==> y + 1 Logical reasoning

More information

A New Look at Generalized Rewriting in Type Theory

A New Look at Generalized Rewriting in Type Theory A New Look at Generalized Rewriting in Type Theory Matthieu Sozeau Harvard University 1st Coq Workshop August 21th 2009 Munich, Germany Generalized Rewriting Equational reasoning x = y - x + 1 ==> y +

More information

0.1 Random useful facts. 0.2 Language Definition

0.1 Random useful facts. 0.2 Language Definition 0.1 Random useful facts Lemma double neg : P : Prop, {P} + { P} P P. Lemma leq dec : n m, {n m} + {n > m}. Lemma lt dec : n m, {n < m} + {n m}. 0.2 Language Definition Definition var := nat. Definition

More information

Well Founded Relations and Recursion

Well Founded Relations and Recursion Well Founded Relations and Recursion Roger Bishop Jones Abstract Fixed points, well founded relations and a recursion theorem. http://www.rbjones.com/rbjpub/pp/doc/t005.pdf Created 2004/07/15 Last Modified

More information

mathlib: Lean s mathematical library

mathlib: Lean s mathematical library mathlib: Lean s mathematical library Johannes Hölzl VU Amsterdam Lean Together 2019 Amsterdam Introduction what is mathlib (classical) mathematical library for Lean classical linear algebra, number theory,

More information

Using the Prover I: Lee Pike. June 3, NASA Langley Formal Methods Group Using the Prover I:

Using the Prover I: Lee Pike. June 3, NASA Langley Formal Methods Group Using the Prover I: Basic Basic NASA Langley Formal Methods Group lee.s.pike@nasa.gov June 3, 2005 Basic Sequents Basic Sequent semantics: The conjunction of the antecedents above the turnstile implies the disjunction of

More information

Verifiable Security of Boneh-Franklin Identity-Based Encryption. Federico Olmedo Gilles Barthe Santiago Zanella Béguelin

Verifiable Security of Boneh-Franklin Identity-Based Encryption. Federico Olmedo Gilles Barthe Santiago Zanella Béguelin Verifiable Security of Boneh-Franklin Identity-Based Encryption Federico Olmedo Gilles Barthe Santiago Zanella Béguelin IMDEA Software Institute, Madrid, Spain 5 th International Conference on Provable

More information

Review. Principles of Programming Languages. Equality. The Diamond Property. The Church-Rosser Theorem. Corollaries. CSE 230: Winter 2007

Review. Principles of Programming Languages. Equality. The Diamond Property. The Church-Rosser Theorem. Corollaries. CSE 230: Winter 2007 CSE 230: Winter 2007 Principles of Programming Languages Lecture 12: The λ-calculus Ranjit Jhala UC San Diego Review The lambda calculus is a calculus of functions: e := x λx. e e 1 e 2 Several evaluation

More information

Red-black Trees of smlnj Studienarbeit. Angelika Kimmig

Red-black Trees of smlnj Studienarbeit. Angelika Kimmig Red-black Trees of smlnj Studienarbeit Angelika Kimmig January 26, 2004 Contents 1 Introduction 1 2 Development of the Isabelle/HOL Theory RBT 2 2.1 Translation of Datatypes and Functions into HOL......

More information

Solutions to Exercises. Solution to Exercise 2.4. Solution to Exercise 2.5. D. Sabel and M. Schmidt-Schauß 1

Solutions to Exercises. Solution to Exercise 2.4. Solution to Exercise 2.5. D. Sabel and M. Schmidt-Schauß 1 D. Sabel and M. Schmidt-Schauß 1 A Solutions to Exercises Solution to Exercise 2.4 We calculate the sets of free and bound variables: FV ((λy.(y x)) (λx.(x y)) (λz.(z x y))) = FV ((λy.(y x)) (λx.(x y)))

More information

Analyse et Conception Formelle. Lesson 4. Proofs with a proof assistant

Analyse et Conception Formelle. Lesson 4. Proofs with a proof assistant Analyse et Conception Formelle Lesson 4 Proofs with a proof assistant T. Genet (ISTIC/IRISA) ACF-4 1 / 26 Prove logic formulas... to prove programs fun nth:: "nat => a list => a" where "nth 0 (x#_)=x"

More information

Principles of Program Analysis: Control Flow Analysis

Principles of Program Analysis: Control Flow Analysis Principles of Program Analysis: Control Flow Analysis Transparencies based on Chapter 3 of the book: Flemming Nielson, Hanne Riis Nielson and Chris Hankin: Principles of Program Analysis. Springer Verlag

More information

Extending the Lambda Calculus: An Eager Functional Language

Extending the Lambda Calculus: An Eager Functional Language Syntax of the basic constructs: Extending the Lambda Calculus: An Eager Functional Language canonical forms z cfm ::= intcfm boolcfm funcfm tuplecfm altcfm intcfm ::= 0 1-1... boolcfm ::= boolconst funcfm

More information

Verifying Probabilistic Programs using the HOL Theorem Prover Joe Hurd p.1/32

Verifying Probabilistic Programs using the HOL Theorem Prover Joe Hurd p.1/32 Verifying Probabilistic Programs using the HOL Theorem Prover Joe Hurd joe.hurd@cl.cam.ac.uk University of Cambridge Verifying Probabilistic Programs using the HOL Theorem Prover Joe Hurd p.1/32 Contents

More information

Gerwin Klein, June Andronick, Ramana Kumar S2/2016

Gerwin Klein, June Andronick, Ramana Kumar S2/2016 COMP4161: Advanced Topics in Software Verification {} Gerwin Klein, June Andronick, Ramana Kumar S2/2016 data61.csiro.au Content Intro & motivation, getting started [1] Foundations & Principles Lambda

More information

M a s t e r r e s e a rc h I n t e r n s h i p. Formalisation of Ground Inference Systems in a Proof Assistant

M a s t e r r e s e a rc h I n t e r n s h i p. Formalisation of Ground Inference Systems in a Proof Assistant M a s t e r r e s e a rc h I n t e r n s h i p Master Thesis Formalisation of Ground Inference Systems in a Proof Assistant Domain: Data Structures and Algorithms - Logic in Computer Science Author: Mathias

More information

Roy L. Crole. Operational Semantics Abstract Machines and Correctness. University of Leicester, UK

Roy L. Crole. Operational Semantics Abstract Machines and Correctness. University of Leicester, UK Midlands Graduate School, University of Birmingham, April 2008 1 Operational Semantics Abstract Machines and Correctness Roy L. Crole University of Leicester, UK Midlands Graduate School, University of

More information

Discrete Mathematics Review

Discrete Mathematics Review CS 1813 Discrete Mathematics Discrete Mathematics Review or Yes, the Final Will Be Comprehensive 1 Truth Tables for Logical Operators P Q P Q False False False P Q False P Q False P Q True P Q True P True

More information

TECHNOLOGY PROJECTS Reflexive Semantic Domains for VDM-SL. Status: Draft Type: Name Location Signature Date. Name Function Signature Date

TECHNOLOGY PROJECTS Reflexive Semantic Domains for VDM-SL. Status: Draft Type: Name Location Signature Date. Name Function Signature Date Title: Status: Draft Type: Author: Name Location Signature Date Roger Bishop Jones Authorisation for Issue: Name Function Signature Date Abstract: Distribution: A construction within of a reflexive domain

More information

Computing Square Roots using the Babylonian Method

Computing Square Roots using the Babylonian Method Computing Square Roots using the Babylonian Method René Thiemann February 16, 2013 Abstract We implement the Babylonian method [1] to compute square roots of numbers. We provide precise algorithms for

More information

Bringing class diagrams to life

Bringing class diagrams to life Bringing class diagrams to life Luis S. Barbosa & Sun Meng DI-CCTC, Minho University, Braga & CWI, Amsterdam UML & FM Workshop 2009 Rio de Janeiro 8 December, 2009 Formal Methods proofs problems structures

More information

1 Buffon s Needle Problem 2

1 Buffon s Needle Problem 2 Buffon s Needle Problem Manuel Eberl August 16, 2018 Abstract In the 18th century, Georges-Louis Leclerc, Comte de Buffon posed and later solved the following problem [1, 2], which is often called the

More information

Programming Languages

Programming Languages CSE 230: Winter 2010 Principles of Programming Languages Lecture 10: Programming in λ-calculusc l l Ranjit Jhala UC San Diego Review The lambda calculus is a calculus of functions: e := x λx. e e 1 e 2

More information

CS 6110 Lecture 21 The Fixed-Point Theorem 8 March 2013 Lecturer: Andrew Myers. 1 Complete partial orders (CPOs) 2 Least fixed points of functions

CS 6110 Lecture 21 The Fixed-Point Theorem 8 March 2013 Lecturer: Andrew Myers. 1 Complete partial orders (CPOs) 2 Least fixed points of functions CS 6110 Lecture 21 The Fixed-Point Theorem 8 March 2013 Lecturer: Andrew Myers We saw that the semantics of the while command are a fixed point. We also saw that intuitively, the semantics are the limit

More information

Theorems for free : a (calculational) introduction

Theorems for free : a (calculational) introduction Theorems for free : a (calculational) introduction J.N. Oliveira Dept. Informática, Universidade do Minho Braga, Portugal 2003 (last update: 2013) Parametric polymorphism by example Function countbits

More information

The Divergence of the Prime Harmonic Series

The Divergence of the Prime Harmonic Series The Divergence of the Prime Harmonic Series Manuel Eberl December 16, 2018 Abstract In this work, we prove the lower bound ln(h n ) ln( 5 3 ) for the partial sum of the Prime Harmonic series and, based

More information

1. The Method of Coalgebra

1. The Method of Coalgebra 1. The Method of Coalgebra Jan Rutten CWI Amsterdam & Radboud University Nijmegen IMS, Singapore - 15 September 2016 Overview of Lecture one 1. Category theory (where coalgebra comes from) 2. Algebras

More information

Relational Algebra by Way of Adjunctions. Jeremy Gibbons (joint work with Fritz Henglein, Ralf Hinze, Nicolas Wu) DBPL, October 2015

Relational Algebra by Way of Adjunctions. Jeremy Gibbons (joint work with Fritz Henglein, Ralf Hinze, Nicolas Wu) DBPL, October 2015 Relational Algebra by Way of Adjunctions Jeremy Gibbons (joint work with Fritz Henglein, Ralf Hinze, Nicolas Wu) DBPL, October 2015 Relational Algebra by Way of Adjunctions 2 1. Summary bulk types (sets,

More information

Math 564 Homework 1. Solutions.

Math 564 Homework 1. Solutions. Math 564 Homework 1. Solutions. Problem 1. Prove Proposition 0.2.2. A guide to this problem: start with the open set S = (a, b), for example. First assume that a >, and show that the number a has the properties

More information

Computer-supported. Modeling and Reasoning. Computer-supported. Exercises and Solutions (Isabelle 2004)

Computer-supported. Modeling and Reasoning. Computer-supported. Exercises and Solutions (Isabelle 2004) Dipl.-Inf. Achim D. Brucker Dr. Burkhart Wolff Computer-supported Modeling and Reasoning http://www.infsec.ethz.ch/ education/permanent/csmr/ (rev. 16826) Computer-supported Modeling and Reasoning Exercises

More information

Iris: Higher-Order Concurrent Separation Logic. Lecture 6: Case Study: foldr

Iris: Higher-Order Concurrent Separation Logic. Lecture 6: Case Study: foldr 1 Iris: Higher-Order Concurrent Separation Logic Lecture 6: Case Study: foldr Lars Birkedal Aarhus University, Denmark November 10, 2017 2 Overview Earlier: Operational Semantics of λ ref,conc e, (h, e)

More information

The Lifting Lemma. Ralf Hinze

The Lifting Lemma. Ralf Hinze The Lifting Lemma Ralf Hinze Computing Laboratory, University of Oxford Wolfson Building, Parks Road, Oxford, OX1 3QD, England ralf.hinze@comlab.ox.ac.uk http://www.comlab.ox.ac.uk/ralf.hinze/ June 2009

More information