Recursive Definitions of Monadic Functions Alexander Krauss le l be a Is λ = β HO L ∀ α → Theorem Proving Group Technische Universität München PAR-10, Edinburgh, July 15, 2010 Definitions in HOL No notion of reduction (only βη , no δι) Definitions in HOL No notion of reduction (only βη , no δι) • System never unfolds definitions Definitions in HOL No notion of reduction (only βη , no δι) • System never unfolds definitions • Strong Normalization: Very weak Definitions in HOL No notion of reduction (only βη , no δι) • System never unfolds definitions • Strong Normalization: Very weak • Unfolding definitions requires explicit rewriting (by simp) Definitions in HOL No notion of reduction (only βη , no δι) • System never unfolds definitions • Strong Normalization: Very weak • Unfolding definitions requires explicit rewriting (by simp) • Some things cannot be reduced at all (e.g. 0 = undefined, where undefined :: α) Definitions in HOL No notion of reduction (only βη , no δι) • System never unfolds definitions • Strong Normalization: Very weak • Unfolding definitions requires explicit rewriting (by simp) • Some things cannot be reduced at all (e.g. 0 = undefined, where undefined :: α) Code Generation (= fast rewriting) Definitions in HOL No notion of reduction (only βη , no δι) • System never unfolds definitions • Strong Normalization: Very weak • Unfolding definitions requires explicit rewriting (by simp) • Some things cannot be reduced at all (e.g. 0 = undefined, where undefined :: α) Code Generation (= fast rewriting) • Equational theorems → Programs (SML, Haskell, OCaml) Definitions in HOL No notion of reduction (only βη , no δι) • System never unfolds definitions • Strong Normalization: Very weak • Unfolding definitions requires explicit rewriting (by simp) • Some things cannot be reduced at all (e.g. 0 = undefined, where undefined :: α) Code Generation (= fast rewriting) • Equational theorems → Programs (SML, Haskell, OCaml) • (+) Any equations, not just definitions Definitions in HOL No notion of reduction (only βη , no δι) • System never unfolds definitions • Strong Normalization: Very weak • Unfolding definitions requires explicit rewriting (by simp) • Some things cannot be reduced at all (e.g. 0 = undefined, where undefined :: α) Code Generation (= fast rewriting) • Equational theorems → Programs (SML, Haskell, OCaml) • (+) Any equations, not just definitions • (–) Partial correctness only (f x = f x) The joy of being classical Definition while :: (α ⇒ bool ) ⇒ (α ⇒ α) ⇒ α ⇒ α while P f x = if (∃n. P (f n (x ))) then f (LEAST n. P (f n (x ))) (x ) else undefined The joy of being classical Definition while :: (α ⇒ bool ) ⇒ (α ⇒ α) ⇒ α ⇒ α while P f x = if (∃n. P (f n (x ))) then f (LEAST n. P (f n (x ))) (x ) else undefined Lemma while P f x = if (P x ) then x else while P f (f x ) The joy of being classical Definition while :: (α ⇒ bool ) ⇒ (α ⇒ α) ⇒ α ⇒ α while P f x = if (∃n. P (f n (x ))) then f (LEAST n. P (f n (x ))) (x ) else undefined Lemma while P f x = if (P x ) then x else while P f (f x ) Proof by case analysis. The joy of being classical Definition while :: (α ⇒ bool ) ⇒ (α ⇒ α) ⇒ α ⇒ α while P f x = if (∃n. P (f n (x ))) then f (LEAST n. P (f n (x ))) (x ) else undefined Lemma [code] while P f x = if (P x ) then x else while P f (f x ) Proof by case analysis. Packages Specification Primitive definition Packages Recursive Equations Specification Primitive definition Packages Recursive Equations Specification Specification Proof Primitive definition Packages Recursive Equations Specification Specification Proof Primitive definition Reasoning tools Packages Recursive Equations Specification Specification Proof Reasoning tools Induction rule Primitive definition Packages Recursive Equations Specification Specification Proof Reasoning tools Induction rule Primitive definition Engineering challenges: • Find suitable constructions • Make sure that exactly the same comes out as is put in • What are “good” reasoning tools? • Do not expose internal operation Packages Recursive Equations Specification Specification Proof Reasoning tools Induction rule Primitive definition Engineering challenges: • Find suitable constructions • Make sure that exactly the same comes out as is put in • What are “good” reasoning tools? • Do not expose internal operation Isabelle’s Function Package AK 2005-2008, extending Slind 1992-1996 Functionality: • Wellfounded (≈ terminating) recursion • Automated termination proofs (lex. orders, size-change, . . . ) • Domain-predicate-based partiality: dom x −→ f x = ... • Nested recursion (dom helps) • Tail recursion (fragile) • Convenience: Pattern matching, Mutual recursion Isabelle’s Function Package AK 2005-2008, extending Slind 1992-1996 Functionality: • Wellfounded (≈ terminating) recursion • Automated termination proofs (lex. orders, size-change, . . . ) • Domain-predicate-based partiality: dom x −→ f x = ... • Nested recursion (dom helps) • Tail recursion (fragile) • Convenience: Pattern matching, Mutual recursion Issues: • Only conditional equations for partial functions • Recursion structure must be analyzable (no arbitrary HO recursion, eg. continuations) • No corecursion • Complex Imperative HOL Ref.alloc :: α ⇒ heap ⇒ α ref × heap Ref.get :: heap ⇒ α ref ⇒ α Ref.set :: α ref ⇒ α ⇒ heap ⇒ heap Imperative HOL Ref.alloc :: α ⇒ heap ⇒ α ref × heap Ref.get :: heap ⇒ α ref ⇒ α Ref.set :: α ref ⇒ α ⇒ heap ⇒ heap datatype α Heap = Heap (heap ⇒ (α × heap) option) where exec (Heap f) = f return :: α ⇒ α Heap »= :: α Heap ⇒ (α ⇒ β Heap) ⇒ β Heap return x = Heap (Some ◦ Pair x) f »= g = Heap (λh. case exec f h of None ⇒ None | Some (x, h 0) ⇒ exec (g x) h 0) Imperative HOL Ref.alloc :: α ⇒ heap ⇒ α ref × heap Ref.get :: heap ⇒ α ref ⇒ α Ref.set :: α ref ⇒ α ⇒ heap ⇒ heap Ref.ref :: α ⇒ α ref Heap Ref.lookup :: α ref ⇒ α Heap Ref.update :: α ref ⇒ α ⇒ unit Heap datatype α Heap = Heap (heap ⇒ (α × heap) option) where exec (Heap f) = f return :: α ⇒ α Heap »= :: α Heap ⇒ (α ⇒ β Heap) ⇒ β Heap return x = Heap (Some ◦ Pair x) f »= g = Heap (λh. case exec f h of None ⇒ None | Some (x, h 0) ⇒ exec (g x) h 0) Imperative HOL Ref.alloc :: α ⇒ heap ⇒ α ref × heap Ref.get :: heap ⇒ α ref ⇒ α Ref.set :: α ref ⇒ α ⇒ heap ⇒ heap Ref.ref :: α ⇒ α ref Heap Ref.lookup :: α ref ⇒ α Heap Ref.update :: α ref ⇒ α ⇒ unit Heap datatype α Heap = Heap (heap ⇒ (α × heap) option) where exec (Heap f) = f return :: α ⇒ α Heap »= :: α Heap ⇒ (α ⇒ β Heap) ⇒ β Heap return x = Heap (Some ◦ Pair x) f »= g = Heap (λh. case exec f h of None ⇒ None | Some (x, h 0) ⇒ exec (g x) h 0) Notation: • !r abbreviates Ref.lookup r etc. • Do-notation: do x ← f; g x done Problematic Recursion datatype α node = Empty | Node α (α node ref) Problematic Recursion datatype α node = Empty | Node α (α node ref) traverse :: α node ⇒ α list Heap traverse Empty = return [ ] traverse (Node x r) = do tl ← !r; xs ← traverse tl; return (x:xs) done Problematic Recursion datatype α node = Empty | Node α (α node ref) traverse :: α node ⇒ α list Heap traverse Empty = return [ ] traverse (Node x r) = do tl ← !r; xs ← traverse tl; return (x:xs) done Recursion structure? Problematic Recursion datatype α node = Empty | Node α (α node ref) traverse :: α node ⇒ α list Heap traverse Empty = return [ ] traverse (Node x r) = do tl ← !r; xs ← traverse tl; return (x:xs) done Recursion structure? Total?? Problematic Recursion datatype α node = Empty | Node α (α node ref) traverse :: α node ⇒ α list Heap traverse Empty = return [ ] traverse (Node x r) = do tl ← !r; xs ← traverse tl; return (x:xs) done Recursion structure? Total?? Induction Rule??? Problematic Recursion datatype α node = Empty | Node α (α node ref) traverse :: α node ⇒ α list Heap traverse Empty = return [ ] traverse (Node x r) = do tl ← !r; xs ← traverse tl; return (x:xs) done Recursion structure? Total?? Induction Rule??? PANIC!!! Problematic Recursion datatype α node = Empty | Node α (α node ref) traverse :: α node ⇒ α list Heap traverse Empty = return [ ] traverse (Node x r) = do tl ← !r; xs ← traverse tl; return (x:xs) done Recursion structure? Total?? Induction Rule??? PANIC!!! But: Such definitions are always consistent. Well-founded recursion is just the wrong tool. Key Ideas Domain theory is the better tool. Key Ideas Domain theory is the better tool. ⊥ can be hidden in the monad. Key Ideas Domain theory is the better tool. ⊥ can be hidden in the monad. Don’t care about termination. Outline HOLCF Option Monad Heap Monad HOLCF Basics Complete partial order v :: α ⇒ α ⇒ bool chain :: (nat ⇒ α) ⇒ bool chain Y ←→ (∀ i. Y i v Y (Suc i)) | :: α set ⇒ α ⇒ bool S | x ←→ (∀ y. y ∈ S −→ y v x) ∧ (∀ u. (∀ y. y ∈ S −→ y v u) −→ x v u) F F :: (nat ⇒ α) ⇒ α ( i. Y i) = (THE x. range Y | x) cont :: (α ⇒ β) ⇒ bool F cont f ←→ (∀ Y. chain Y −→ range (λi. f (Y i)) | f ( i. Y i)) Fixed-Point operator: cont F −→ fixp F = F (fixp F) Option pcpo/monad datatype α option = None | Some α Pcpo structure: x v y = (x = None ∨ x = y) (thus ⊥ = None) Option pcpo/monad datatype α option = None | Some α Pcpo structure: x v y = (x = None ∨ x = y) Monad structure: return x = Some x None »= f = None Some y »= f = f y (thus ⊥ = None) Option pcpo/monad datatype α option = None | Some α Pcpo structure: x v y = (x = None ∨ x = y) (thus ⊥ = None) Monad structure: return x = Some x None »= f = None Some y »= f = f y Classical: Abuse None as result of nonterminating computation. Option pcpo/monad datatype α option = None | Some α Pcpo structure: x v y = (x = None ∨ x = y) (thus ⊥ = None) Monad structure: return x = Some x None »= f = None Some y »= f = f y Classical: Abuse None as result of nonterminating computation. Legal, but not a program: if f x = None then A else B Steps To define f with f x = F f x: Steps To define f with f x = F f x: 1. Prove that F is continuous. Steps To define f with f x = F f x: 1. Prove that F is continuous. 2. Define f = fixp F. Steps To define f with f x = F f x: 1. Prove that F is continuous. 2. Define f = fixp F. 3. Conclude the equation f x = F f x using the fixed-point theorem. Example function In ML: fun trace n = if n = 0 then [] else if even n then n :: trace (step n) else trace (step n) Example function In ML: fun trace n = if n = 0 then [] else if even n then n :: trace (step n) else trace (step n) In monadic HOL: trace n = (if n = 0 then return [ ] else do tl ← trace (step n); (if even n then return (n:tl) else return tl) done) Continuity Proofs (L AM ) (B IND ) (C ONST ) (R EC ) (I F ) (∀ y. cont (λx. f x y)) −→ cont f cont f −→ (∀ y. cont (g y)) −→ cont (λx. f x »= (λy. g y x)) cont (λf. c) cont (λf. f x) cont f −→ cont g −→ cont (λx. if b then f x else g x) Continuity Proofs (L AM ) (B IND ) (C ONST ) (R EC ) (I F ) (∀ y. cont (λx. f x y)) −→ cont f cont f −→ (∀ y. cont (g y)) −→ cont (λx. f x »= (λy. g y x)) cont (λf. c) cont (λf. f x) cont f −→ cont g −→ cont (λx. if b then f x else g x) 1. cont (λtrace n. if n = 0 then return [ ] else do tl ← trace (step n); (if even n then return (n:tl) else return tl) done) Continuity Proofs (L AM ) (B IND ) (C ONST ) (R EC ) (I F ) (∀ y. cont (λx. f x y)) −→ cont f cont f −→ (∀ y. cont (g y)) −→ cont (λx. f x »= (λy. g y x)) cont (λf. c) cont (λf. f x) cont f −→ cont g −→ cont (λx. if b then f x else g x) 1. ∀ n. cont (λtrace. if n = 0 then return [ ] else do tl ← trace (step n); (if even n then return (n:tl) else return tl) done) Continuity Proofs (L AM ) (B IND ) (C ONST ) (R EC ) (I F ) (∀ y. cont (λx. f x y)) −→ cont f cont f −→ (∀ y. cont (g y)) −→ cont (λx. f x »= (λy. g y x)) cont (λf. c) cont (λf. f x) cont f −→ cont g −→ cont (λx. if b then f x else g x) 1. ∀ n. cont (λtrace. return [ ]) 2. ∀ n. cont (λtrace. do tl ← trace (step n); (if even n then return (n:tl) else return tl) done) Continuity Proofs (L AM ) (B IND ) (C ONST ) (R EC ) (I F ) (∀ y. cont (λx. f x y)) −→ cont f cont f −→ (∀ y. cont (g y)) −→ cont (λx. f x »= (λy. g y x)) cont (λf. c) cont (λf. f x) cont f −→ cont g −→ cont (λx. if b then f x else g x) 1. ∀ n. cont (λtrace. do tl ← trace (step n); (if even n then return (n:tl) else return tl) done) Continuity Proofs (L AM ) (B IND ) (C ONST ) (R EC ) (I F ) (∀ y. cont (λx. f x y)) −→ cont f cont f −→ (∀ y. cont (g y)) −→ cont (λx. f x »= (λy. g y x)) cont (λf. c) cont (λf. f x) cont f −→ cont g −→ cont (λx. if b then f x else g x) 1. ∀ n. cont (λtrace. trace (step n)) 2. ∀ n tl. cont (λtrace. if even n then return (n:tl) else return tl) Continuity Proofs (L AM ) (B IND ) (C ONST ) (R EC ) (I F ) (∀ y. cont (λx. f x y)) −→ cont f cont f −→ (∀ y. cont (g y)) −→ cont (λx. f x »= (λy. g y x)) cont (λf. c) cont (λf. f x) cont f −→ cont g −→ cont (λx. if b then f x else g x) 1. ∀ n tl. cont (λtrace. if even n then return (n:tl) else return tl) Continuity Proofs (L AM ) (B IND ) (C ONST ) (R EC ) (I F ) (∀ y. cont (λx. f x y)) −→ cont f cont f −→ (∀ y. cont (g y)) −→ cont (λx. f x »= (λy. g y x)) cont (λf. c) cont (λf. f x) cont f −→ cont g −→ cont (λx. if b then f x else g x) No subgoals! “Continuous by construction” Induction Rule Abstract fixed-point induction (from HOLCF): adm P cont F P⊥ ∀ f. P f −→ P (F f) P (fixp F) where adm P = (∀ Y. chain Y −→ (∀ i. P (Y i)) −→ P ( F i. Y i)) Induction Rule Abstract fixed-point induction (from HOLCF): adm P cont F P⊥ ∀ f. P f −→ P (F f) P (fixp F) where adm P = (∀ Y. chain Y −→ (∀ i. P (Y i)) −→ P ( F i. Y i)) Instantiate P with λf. ∀ x y. f x = Some y −→ Q x y. Concrete fixed-point induction: ∀ f x y. (∀ z r. f z = Some r −→ Q z r) −→ F f x = Some y −→ Q x y f x = Some y −→ Q x y Induction Rule Refinement ∀ trace n ys. (∀ z r. trace z = Some r −→ Q z r) −→ (if n = 0 then return [ ] else do tl ← trace (step n); (if even n then return (n:tl) else return tl) done) = Some ys −→ Q n ys trace n = Some ys −→ Q n ys Induction Rule Refinement ∀ trace n ys. (∀ z r. trace z = Some r −→ Q z r) −→ (if n = 0 then return [ ] else do tl ← trace (step n); (if even n then return (n:tl) else return tl) done) = Some ys −→ Q n ys trace n = Some ys −→ Q n ys Q 0 [] ∀ n tl. n 6= 0 −→ Q (step n) tl −→ even n −→ Q n (n:tl) ∀ n tl. n 6= 0 −→ Q (step n) tl −→ ¬ even n −→ Q n tl trace n = Some ys −→ Q n ys Induction Rule Refinement Refinement steps: 1. Decompose program structure: I (t »= (λx. f x)) = Some y becomes t = Some x and f x = Some y I Some t = Some y becomes t = y. I Split conditionals like (if b then x else x 0) = Some y Induction Rule Refinement Refinement steps: 1. Decompose program structure: I (t »= (λx. f x)) = Some y becomes t = Some x and f x = Some y I Some t = Some y becomes t = y. I Split conditionals like (if b then x else x 0) = Some y 2. Use induction hypothesis to replace f z = Some r by Q z r. Induction Rule Refinement Refinement steps: 1. Decompose program structure: I (t »= (λx. f x)) = Some y becomes t = Some x and f x = Some y I Some t = Some y becomes t = y. I Split conditionals like (if b then x else x 0) = Some y 2. Use induction hypothesis to replace f z = Some r by Q z r. 3. Clean up context: Substitute v = t. Heap pcpo/monad datatype α Heap = Heap.Heap (heap * α × heap) Pcpo structure: Heap f v Heap g ←→ (∀ h. execute (Heap f) h v execute (Heap g) h) Heap pcpo/monad datatype α Heap = Heap.Heap (heap * α × heap) Pcpo structure: Heap f v Heap g ←→ (∀ h. execute (Heap f) h v execute (Heap g) h) Tedious but straightforward: »= is still continuous. Recursion in the Heap Monad Proceed as before: 1. Prove that F is continuous (as before) 2. Define function f = fixp F (as before) 3. Equation follows by fixed-point theorem (as before) Recursion in the Heap Monad Proceed as before: 1. Prove that F is continuous (as before) 2. Define function f = fixp F (as before) 3. Equation follows by fixed-point theorem (as before) What is different: • Induction Rule: f x = Some y is replaced by (h, h 0, y) ∈ [[ f x ]] (must prove admissibility) • Refinement includes steps for Imperative HOL primitives. E.g., (h, h 0, y) ∈ [[ !r ]] is replaced by y = Ref.get h r and h 0 = h. Example: Imperative Unification datatype α rtrm = Var α (α rtrm ref option) | Const α | App (α rtrm ref) (α rtrm ref) occurs r1 r2 = do t ← !r2 ; (case t of Var n σ ⇒ if r1 = r2 then return True else case σ of None ⇒ return False | Some r 0 ⇒ occurs r1 r 0 | Const n ⇒ return False | App r3 r4 ⇒ do b ← occurs r1 r3 ; (if b then return True else occurs r1 r4 ) done) done Example: Imperative Unification ∀ r1 h n σ. Ref.get h r1 = Var n σ −→ P r1 r1 h h True ∀ r1 r2 h n. r1 6= r2 −→ Ref.get h r2 = Var n None −→ P r1 r2 h h False ∀ r1 r2 h 0 y h n r 0. r1 6= r2 −→ P r1 r 0 h h 0 y −→ Ref.get h r2 = Var n (Some r 0) −→ P r1 r2 h h 0 y ∀ r1 r2 h n. Ref.get h r2 = Const n −→ P r1 r2 h h False ∀ r1 r2 h 0 h r3 r4 b. Ref.get h r2 = App r3 r4 −→ P r1 r3 h h 0 b −→ b −→ P r1 r2 h h 0 True ∀ r1 r2 h 00 y h r3 r4 h 0 b. Ref.get h r2 = App r3 r4 −→ P r1 r3 h h 0 b −→ ¬ b −→ P r1 r4 h 0 h 00 y −→ P r1 r2 h h 00 y (y, r1 , r2 ) ∈ [[ occurs h h 0 ]] −→ P h h 0 y r1 r2 Other Monads? Heap Imperative programs Other Monads? Heap Option/Error Imperative programs Ordinary partial functions Other Monads? Heap Option/Error Identity Imperative programs Ordinary partial functions Tail recursion (no »=) Other Monads? Heap Option/Error Identity Set Imperative programs Ordinary partial functions Tail recursion (no »=) Nothing new (inductive sets) Other Monads? Heap Option/Error Identity Set Imperative programs Ordinary partial functions Tail recursion (no »=) Nothing new (inductive sets) In general??? No idea :-}
© Copyright 2025 Paperzz