A Gentle Introduction to F*

fstar-logo

Danel Ahman, Inria Paris

EUTypes Summer School

Ohrid, Macedonia, 10 August, 2018

Schedule

  • Today: A Gentle Introduction to F* (Purely Functional Programs)

  • Tomorrow: Verifying Stateful Programs in F*

  • Sunday: Monotonic State in F*

  • Sunday: F*'s Extensible Effect System and Metaprogramming in F*

Slides, code, exercises, and setup instructions
@
https://danelahman.github.io/teaching/eutypes2018/


Please ask questions at any time!

Program verification: Shall the twain ever meet?

 

Interactive proof assistants Semi-automated verifiers of
imperative programs
    Coq,   air         Dafny,
    Agda,       FramaC,
    Lean,   gap         Why3,
    Isabelle       Liquid Types
  • Left corner: very expressive higher-order logics, interactive proving,
    tactics, but mostly only purely functional programming

  • Right corner: effectful programming, SMT-based automation,
    but only very weak logics

Bridging the gap: F*

  • Functional programming language with effects

    • like OCaml, F#, Haskell,
      let incr (r:ref a) = r := !r + 1
    • F* extracted to OCaml or F# by default
    • subset of F* extracted to efficient C code (Low* and KreMLin)
  • Semi-automated verification system using SMT

    • like Dafny, FramaC, Why3, Liquid Types,
  • Interactive proof assistant based on dependent types

    • like Coq, Agda, Lean,

F* in action, at scale

  • Functional programming language with effects

    • F* is programmed in F*, but not (yet) verified
  • Semi-automated verification system

    • Project Everest: verify and deploy new, efficient HTTPS stack
      • miTLS*: Verified reference implementation of TLS (1.2 and 1.3)
      • HACL*: High-Assurance Cryptographic Library
      • Vale: Verified Assembly Language for Everest
  • Proof assistant based on dependent types

    • Fallback when SMT fails; also for mechanized metatheory
      • MicroFStar: Fragment of F* formalized in F*
      • Wys*: Verified DSL for secure multi-party computations
      • ReVer: Verified compiler to reversible circuits
    • F*'s new metaprogramming framework increasingly used in Everest

The current F* team

Microsoft Research (US, UK, India), Inria Paris, MIT, Rosario, …

Danel Ahman
Benjamin Beurdouche
Karthikeyan Bhargavan
Barry Bond
Antoine Delignat-Lavaud   
Victor Dumitrescu
Cédric Fournet
Chris Hawblitzel
Cătălin Hriţcu
Markulf Kohlweiss
Qunyan Mangus
Kenji Maillard
Asher Manning
Guido Martínez
Zoe Paraskevopoulou
Clément Pit-Claudel
Jonathan Protzenko
Tahina Ramananandro
Aseem Rastogi
Nikhil Swamy (benevolent dictator)
Christoph M. Wintersteiger
Santiago Zanella-Béguelin
Gustavo Varo

The rest of this lecture

  • The functional core of F*

  • Verifying purely functional programs

  • Using very simple examples throughout

  • Small hands-on exercises in the end (for the exercise classes)

The functional core of F*

  • Recursive functions

    val factorial : nat -> nat
    
    let rec factorial n =
      if n = 0 then 1 else n * (factorial (n - 1))
  • (Simple) inductive datatypes and pattern matching

    type list (a:Type) =
      | Nil  : list a
      | Cons : hd:a -> tl:list a -> list a
    
    val map : ('a -> 'b) -> list 'a -> list 'b
    
    let rec map f x =
      match x with
      | []     -> []
      | h :: t -> f h :: map f t
  • Lambdas

    map (fun x -> x + 42) [1;2;3]

Refinement types

type nat = x:int{x >= 0}
  • Refinements introduced by type annotations (code unchanged)

    val factorial : nat -> nat
    let rec factorial n = if n = 0 then 1 else n * (factorial (n - 1))
  • Logical obligations discharged by SMT (for else branch, simplified)

    n >= 0, n <> 0 |= n - 1 >= 0
    n >= 0, n <> 0, (factorial (n - 1)) >= 0 |= n * (factorial (n - 1)) >= 0
  • Refinements eliminated by subtyping: nat <: int

    let i : int = factorial 42
    
    let f : x:nat{x > 0} -> int = factorial
  • Different kinds of extensional equality (=, ==, ===)

Dependent types

  • Dependent function types ($\Pi$), here together with refinements:

    val incr : x:int -> y:int{x < y}
    let incr x = x + 1
  • Can express pre- and postconditions of pure functions

    val incr' : x:nat{odd x} -> y:nat{even y}
  • (Parameterised and indexed) inductive datatypes; implicit arguments

    type vec (a:Type) : nat -> Type =
      | Nil  : vec a 0
      | Cons : #n:nat -> hd:a -> tl:vec a n -> vec a (n + 1)
    
    val map : #n:nat -> #a:Type -> #b:Type -> (a -> b) -> vec a n -> vec b n
    let rec map #n #a #b f as =
      match as with
      | Nil        -> Nil
      | Cons hd tl -> Cons (f hd) (map f tl)
    
    val lookup : #a:Type -> #n:nat -> vec a n -> m:nat -> m `less_than` n -> a
    let rec lookup #a #n v m p = ...

Inductive families + refinement types

  • As in e.g. Coq, we could define on the last slide

    type vec (a:Type) : nat -> Type =
      | Nil  : vec a 0
      | Cons : #n:nat -> hd:a -> tl:vec a n -> vec a (n + 1)
    
    val lookup : #a:Type -> #n:nat -> vec a n -> m:nat -> m `less_than` n -> a
    let rec lookup #a #n v m p = ...
  • But can also combine vec with refinement types for more convenience

    val lookup : #a:Type -> #n:nat -> vec a n -> m:nat{m `less_than` n} -> a
    let rec lookup #a #n v m = ...
  • Or we could even just use lists + refinement types

    type list (a:Type) = | Nil : list a | Cons : hd:a -> tl:list a -> list a
    
    val length : #a:Type -> list a -> nat
    let rec length #a l = match l with | Nil -> 0 | Cons _ tl -> 1 + length tl
      
    val lookup : #a:Type -> l:list a -> m:nat{m `less_than` (length l)} -> a
    let rec lookup #a l m = ...

Total functions in F*

  • The F* functions we saw so far were all total

  • Tot effect (default) = no side-effects, terminates on all inputs

    val factorial : nat -> Tot nat
    let rec factorial n =
      if n = 0 then 1 else n * (factorial (n - 1))
  • Quiz: How about giving this weaker type to factorial?

    val factorial : int -> Tot int
  let rec factorial n = if n = 0 then 1 else n * (factorial (n - 1))
                                                             ^^^^^
  Subtyping check failed; expected type (x:int{(x << n)}); got type int

factorial (-1) loops!    (int type in F* is unbounded)

Semantic termination checking

  • based on well-founded ordering on expressions (<<)
    • naturals related by < (negative integers unrelated)
    • inductives related by subterm ordering
    • lex tuples %[a;b;c] with lexicographic ordering
  • order constraints discharged by the SMT solver
  • arbitrary total expression as decreases metric
    val ackermann: m:nat -> n:nat -> Tot nat (decreases %[m;n])
    let rec ackermann m n =
      if m = 0 then n + 1
      else if n = 0 then ackermann (m - 1) 1
      else ackermann (m - 1) (ackermann m (n - 1))
  • default metric is lex ordering of all (non-function) args
    val ackermann: m:nat -> n:nat -> Tot nat

The divergence effect (Dv)

  • We might not want to prove all code terminating

    val factorial : int -> Dv int
  • Some useful code really is not always terminating

    • evaluator for lambda terms
      val eval : exp -> Dv exp
      let rec eval e = 
        match e with
        | App (Lam x e1) e2 -> eval (subst x e2 e1)
        | App e1 e2         -> eval (App (eval e1) e2)
        | Lam x e1          -> Lam x (eval e1)
        | _                 -> e
      
      let main () = eval (App (Lam 0 (App (Var 0) (Var 0)))
                              (Lam 0 (App (Var 0) (Var 0))))
      ./Divergence.exe
    • web servers

F* effect system encapsulates effectful code (Tot and Dv)

  • Pure code cannot call potentially divergent code

  • Only pure code can appear in specifications

    val factorial : int -> Dv int
    
    type tau = x:int{x = factorial (-1)}
    type tau = x:int{x = factorial (-1)}
                     ^^^^^^^^^^^^^^^^^^
    Expected a pure expression; got an expression ... with effect "DIV"
  • Sub-effecting: Tot t <: Dv t

  • So, divergent code can include pure code

    incr 2 + factorial (-1) : Dv int

F* effect system encapsulates effectful code (Tot and GTot)

  • Ghost effect for code used only in specifications

    val sel : #a:Type -> heap -> ref a -> GTot a
  • Sub-effecting: Tot t <: GTot t

  • BUT NOT: GTot t <: Tot t (holds for non-informative types)

  • So, (informative) ghost code cannot be used in total functions

    let f (g:unit -> GTot nat) : Tot (n:nat{n = g ()}) = g ()
    Computed type "n:nat{n = g ()}" and effect "GTot"
    is not compatible with the annotated type "n:nat{n = g ()}" effect "Tot"
  • But total functions can appear in ghost code (regardless of their type)

    let f (g:unit -> Tot nat) : GTot (n:nat{n = g ()}) = g ()

Verifying pure programs

Variant #1: intrinsically (at definition time)

  • Using refinement types (saw this already)
    val factorial : nat -> Tot nat              (* type nat = x:int{x >= 0} *)
  • Can equivalently use pre- and postconditions for this
    val factorial : x:int -> Pure int (requires (x >= 0))
                                      (ensures  (fun y -> y >= 0))
  • Each F* computation type is of the form
    • effect (e.g. Pure)       result type (e.g. int)       spec. (e.g. pre and post)

  • Tot is essentially just an abbreviation
    Tot t = Pure t (requires True) (ensures (fun _ -> True))

Verifying potentially divergent programs

The only variant: intrinsically (partial correctness)

  • Using refinement types
    val factorial : nat -> Dv nat
  • Or the Div computation type (pre- and postconditions)
    val eval_closed : e:exp -> Div exp
                                   (requires (closed e))
                                   (ensures  (fun e' -> Lam? e' /\ closed e'))
    let rec eval_closed e =
      match e with           (* notice there is no match case for variables *)
      | App e1 e2 ->
          let Lam e1' = eval_closed e1 in
          below_subst_beta 0 e1' e2;
          eval_closed (subst (sub_beta e2) e1')
      | Lam e1 -> Lam e1
  • Dv is also just an abbreviation
    Dv t = Div t (requires True) (ensures (fun _ -> True))

Another way to look at this

  • Two classes of types
    • Value types (t): int, list int,
    • Computation types (C): Tot t, Dv t, GTot t
  • Dependent (effectful) function types of the form: x:t -> C

    • argument can't have side-effects, so a value type

  • Two forms of refinement types

    • Refined value types:
      • x:t{p}
    • Refined computation types:
      • Pure t pre post
      • Div t pre post
      • Ghost t pre post
      • these will get more interesting for more interesting effects

Verifying pure programs

Variant #2: extrinsically using SMT-backed lemmas

let rec append (#a:Type) (xs ys : list a) : Tot (list a) = match xs with
  | []       -> ys
  | x :: xs' -> x :: append xs' ys
let rec append_length (#a:Type) (xs ys : list a) 
  : Pure unit
      (requires True)
      (ensures  (fun _ -> length (append xs ys) = length xs + length ys))
      
= match xs with
  | []       -> ()          
 (* nil-VC:  postcondition of () ==> len (app [] ys) = len [] + len ys   *)
 (* nil-VC': True ==> len ys = 0 + len ys                                *)
 
  | x :: xs' -> append_length xs' ys
 (* recursive call's postcondition: len (app xs' ys) = len xs' + len ys  *)
 (* cons-VC:  rec_post ==> len (app (x::xs') ys) = len (x::xs') + len ys *)                
 (* cons-VC': rec_post ==> 1 + len (app xs' ys) = (1 + len xs') + len ys *)
  • Convenient syntactic sugar: the Lemma effect
    Lemma (property) = Pure unit (requires True) (ensures (fun _ -> property))
  • F* also provides a requires-ensures variant of the Lemma effect

Often lemmas are unavoidable

let snoc l h = l @ [h]

val rev : #a:Type -> list a -> Tot (list a)

let rec rev (#a:Type) l =
  match l with
  | []     -> []
  | hd::tl -> snoc (rev tl) hd
val rev_snoc : #a:Type -> l:list a -> h:a ->
                                        Lemma (rev (snoc l h) == h::rev l)


let rec rev_snoc (#a:Type) l h =
  match l with
  | []     -> ()
  | hd::tl -> rev_snoc tl h
val rev_involutive : #a:Type -> l:list a -> Lemma (rev (rev l) == l)

let rec rev_involutive (#a:Type) l =
  match l with
  | []     -> ()
  | hd::tl -> rev_involutive tl; rev_snoc (rev tl) hd

Often lemmas are unavoidable (but SMT can help)

let snoc l h = l @ [h]

val rev : #a:Type -> list a -> Tot (list a)

let rec rev (#a:Type) l =
  match l with
  | []     -> []
  | hd::tl -> snoc (rev tl) hd
val rev_snoc : #a:Type -> l:list a -> h:a -> 
                                        Lemma (rev (snoc l h) == h::rev l)
                                        [SMTPat (rev (snoc l h))]
                                            
let rec rev_snoc (#a:Type) l h =
  match l with
  | []     -> ()
  | hd::tl -> rev_snoc tl h
val rev_involutive : #a:Type -> l:list a -> Lemma (rev (rev l) == l)

let rec rev_involutive (#a:Type) l =
  match l with
  | []     -> ()
  | hd::tl -> rev_involutive tl (*; rev_snoc (rev tl) hd*)

Verifying pure programs

Variant #3: using proof terms


val progress : #e:exp -> #t:typ -> h:typing empty e t ->
                         Pure (cexists (fun e' -> step e e'))
                              (requires (~ (is_value e)))
                              (ensures  (fun _ -> True)) (decreases h)
let rec progress #e #t h =
  match h with
  | TyApp #g #e1 #e2 #t11 #t12 h1 h2 ->
     match e1 with
     | ELam t e1' -> ExIntro (subst (sub_beta e2) e1') (SBeta t e1' e2)
     | _          -> let ExIntro e1' h1' = progress h1 in
                     ExIntro (EApp e1' e2) (SApp1 e2 h1')
  • Note: match exhaustiveness check also semantic (via SMT)

Functional core of F*

  • Variant of dependent type theory

    • $\lambda$, $\Pi$, inductives, matches, universe polymorphism
  • Recursion and semantic termination check

    • potential non-termination is an effect
  • Refinements

    • Refined value types:
      • x:t{p}
    • Refined computation types:
      • Pure t pre post
      • Div t pre post
    • computationally and proof irrelevant, discharged by SMT
  • Subtyping and sub-effecting (<:)

  • Extensional equality (=, ==, ===)

Exercises

Exercise 1: Summing: 0 + 1 + 2 + 3 + …

module Sum

open FStar.Mul

let rec sum_rec (n:nat) = if n > 0 then n + sum_rec (n - 1) else 0

let sum_tot (n:nat) : nat = ((n + 1) * n) / 2

let rec sum_rec_correct (n:nat) : Lemma (sum_rec n = sum_tot n) = 
  
  admit()                       (* replace this admit with a real proof *)

Exercise2: Simply typed stacks (the interface)

Stack.fsti

val stack : Type0  (* type of stacks *)

val empty : stack
val is_empty : stack -> GTot bool

val push : int -> stack -> stack
val pop : stack -> option stack
val top : stack -> option int
val lemma_empty_is_empty : unit -> Lemma (is_empty (empty))

val lemma_push_is_empty : s:stack -> i:int -> 
                                            Lemma (~(is_empty (push i s))) 

val lemma_is_empty_top_some : s:stack{~(is_empty s)} -> 
                                                     Lemma (Some? (top s)) 

val lemma_is_empty_pop_some : s:stack{~(is_empty s)} -> 
                                                     Lemma (Some? (pop s)) 


(* Hint1: You will need to provide some more lemmas about pop and top *)

(* Hint2: You will need to annotate some lemmas with [SMTPat (...)]s *)

Exercise 2: Simply typed stacks (the implementation)

Stack.fst

module Stack

let stack = list int

(* replace these admits with real code and proofs to match Stack.fsti *)

let empty = admit () 
let is_empty = admit ()

let push = admit ()
let pop = admit ()
let top = admit ()

let lemma_empty_is_empty = admit ()
let lemma_push_is_empty = admit ()
let lemma_is_empty_top_some = admit ()
let lemma_is_empty_pop_some = admit ()

Exercise 2: Simply typed stacks (the client)

StackClient.fst

module StackClient

open Stack

[@fail]      (* remove this attribute once you have completed Stack.fst *)
let main() =

  let s0 = empty (* <: stack *) in
    
  lemma_empty_is_empty ();
  assert (is_empty s0);
    
  let s1 = push 3 s0 (* <: stack *) in
  assert (~(is_empty s1));
    
  let s2 = push 4 s1 (* <: stack *) in
  assert (~(is_empty s2));
    
  let i = top s2 (* <: option int *) in
  assert (Some?.v i = 4);
    
  let s3 = pop s2 (* <: option stack *) in
  assert (Some?.v s3 == s1)

Exercise 3: Refinement-typed stacks

RefinedStack.fsti

module RefinedStack

val stack : Type0  (* type of stacks *)

(* Exercise: modify and implement this interface of refined stacks; 
             pop and top must not return in the option type here    *)

(* Hint: compared to Stack.fsti and Stack.fst, you will need to 
         refine stack types below with the is_empty predicate   *)

val empty : stack
val is_empty : stack -> GTot bool

val push : int -> stack -> stack
val pop : stack -> stack           (* before the type was `option stack`*)
val top : stack -> int             (* before the type was `option int`  *)
  • Goal: type a variant of StackClient.fst (now without Some?.v's)
    let main () =
      ...
      let s3 = pop s2 (* <: stack *) in
      assert (s3 == s1)

Next steps in this course

  • Today: A Gentle Introduction to F* (Purely Functional Programs)

  • Tomorrow: Verifying Stateful Programs in F*

  • Sunday: Monotonic State in F*

  • Sunday: F*'s Extensible Effect System and Metaprogramming in F*