Operational Semantics

Similar documents
Denotational semantics

On the coinductive nature of centralizers

1 Introduction. 2 Recap The Typed λ-calculus λ. 3 Simple Data Structures

COSE312: Compilers. Lecture 2 Lexical Analysis (1)

CS:4330 Theory of Computation Spring Regular Languages. Finite Automata and Regular Expressions. Haniel Barbosa

The Trace Partitioning Abstract Domain

CS422 - Programming Language Design

Formal Techniques for Software Engineering: Denotational Semantics

Equivalence of Regular Expressions and FSMs

Register machines L2 18

Course Runtime Verification

Part I: Definitions and Properties

Chapter 5. Finite Automata

Automata-based Verification - III

Closure under the Regular Operations

Undecidable Problems. Z. Sawa (TU Ostrava) Introd. to Theoretical Computer Science May 12, / 65

Theory of computation: initial remarks (Chapter 11)

Theory of Computing Tamás Herendi

Dynamic Noninterference Analysis Using Context Sensitive Static Analyses. Gurvan Le Guernic July 14, 2007

Semantics and Verification of Software

EDA045F: Program Analysis LECTURE 10: TYPES 1. Christoph Reichenbach

(a) Definition of TMs. First Problem of URMs

Lambda-Calculus (I) 2nd Asian-Pacific Summer School on Formal Methods Tsinghua University, August 23, 2010

CSE 200 Lecture Notes Turing machine vs. RAM machine vs. circuits

G54FOP: Lecture 17 & 18 Denotational Semantics and Domain Theory III & IV

Towards a formal language for systemic requirements

PSL Model Checking and Run-time Verification via Testers

Non-emptiness Testing for TMs

Theoretical Computer Science

Denotational Semantics of Programs. : SimpleExp N.

Automata-based Verification - III

Coinductive big-step operational semantics

Language Stability and Stabilizability of Discrete Event Dynamical Systems 1

Models of computation

A call-by-name lambda-calculus machine

Dynamic Semantics. Dynamic Semantics. Operational Semantics Axiomatic Semantics Denotational Semantic. Operational Semantics

The Expressivity of Universal Timed CCP: Undecidability of Monadic FLTL and Closure Operators for Security

Safety Analysis versus Type Inference

Automata: a short introduction

1 Definition of a Turing machine

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

Theory of Computation

An Abstract Interpretation Framework for Semantics and Diagnosis of Lazy Functional-Logic Languages

X-machines - a computational model framework.

Theory of Computation Lecture Notes. Problems and Algorithms. Class Information

Notes on Monoids and Automata

Uses of finite automata

Finite Automata. Mahesh Viswanathan

Lexical Analysis. Reinhard Wilhelm, Sebastian Hack, Mooly Sagiv Saarland University, Tel Aviv University.

Axiomatic Semantics. Stansifer Ch 2.4, Ch. 9 Winskel Ch.6 Slonneger and Kurtz Ch. 11 CSE

Timo Latvala. March 7, 2004

Lecture Notes On THEORY OF COMPUTATION MODULE -1 UNIT - 2

Static Program Analysis

Denotational Semantics

EECS 144/244: Fundamental Algorithms for System Modeling, Analysis, and Optimization

CISC 4090: Theory of Computation Chapter 1 Regular Languages. Section 1.1: Finite Automata. What is a computer? Finite automata

Discrete Fixpoint Approximation Methods in Program Static Analysis

Lecture 17: Language Recognition

Theory of computation: initial remarks (Chapter 11)

Compilers. Lexical analysis. Yannis Smaragdakis, U. Athens (original slides by Sam

CHAPTER 11. A Revision. 1. The Computers and Numbers therein

Linear-time Temporal Logic

CS243, Logic and Computation Nondeterministic finite automata

What is this course about?

Temporal Logic. Stavros Tripakis University of California, Berkeley. We have designed a system. We want to check that it is correct.

arxiv: v2 [cs.fl] 29 Nov 2013

Multicore Semantics and Programming

Cours M.2-6 «Interprétation abstraite: applications à la vérification et à l analyse statique» Examen partiel. Patrick Cousot.

Theory of Computation

Classical Program Logics: Hoare Logic, Weakest Liberal Preconditions

Limitations of OCAML records

ω-automata Automata that accept (or reject) words of infinite length. Languages of infinite words appear:

Typing λ-terms. Types. Typed λ-terms. Base Types. The Typing Relation. Advanced Formal Methods. Lecture 3: Simply Typed Lambda calculus

Variable Automata over Infinite Alphabets

Decentralized Control of Discrete Event Systems with Bounded or Unbounded Delay Communication

Pointer Programs and Undirected Reachability

Theoretical Foundations of the UML

Automata theory. An algorithmic approach. Lecture Notes. Javier Esparza

Proofs, Strings, and Finite Automata. CS154 Chris Pollett Feb 5, 2007.

Temporal Logic Model Checking

Unranked Tree Automata with Sibling Equalities and Disequalities

ECE290 Fall 2012 Lecture 22. Dr. Zbigniew Kalbarczyk

COMPUTER SCIENCE TRIPOS

Let us first give some intuitive idea about a state of a system and state transitions before describing finite automata.

Intelligent Agents. Formal Characteristics of Planning. Ute Schmid. Cognitive Systems, Applied Computer Science, Bamberg University

Simply Typed Lambda Calculus

Sri vidya college of engineering and technology

Chapter 3: Linear temporal logic

Tasks of lexer. CISC 5920: Compiler Construction Chapter 2 Lexical Analysis. Tokens and lexemes. Buffering

Bi-inductive Structural Semantics

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

6.8 The Post Correspondence Problem

Automata on linear orderings

Benefits of Interval Temporal Logic for Specification of Concurrent Systems

Computational Theory

Temporal logics and explicit-state model checking. Pierre Wolper Université de Liège

CSE 505, Fall 2005, Midterm Examination 8 November Please do not turn the page until everyone is ready.

Partial model checking via abstract interpretation

Failure Diagnosis of Discrete Event Systems With Linear-Time Temporal Logic Specifications

Introduction to Turing Machines

Transcription:

Operational Semantics Semantics and applications to verification Xavier Rival École Normale Supérieure Xavier Rival Operational Semantics 1 / 50

Program of this first lecture Operational semantics Mathematical description of the executions of a program 1 a model of programs: transition systems definition, a small step semantics a few common examples 2 trace semantics: a kind of big step semantics finite and infinite executions fixpoint-based definitions notion of compositional semantics Xavier Rival Operational Semantics 2 / 50

Transition systems and small step semantics Definition and properties Outline 1 Transition systems and small step semantics Definition and properties Examples 2 Traces semantics 3 Summary Xavier Rival Operational Semantics 3 / 50

Transition systems and small step semantics Definition and properties Definition We will characterize a program by: states: photography of the program status at an instant of the execution execution steps: how do we move from one state to the next one Definition: transition systems (TS) A transition system is a tuple (S, ) where: Note: S is the set of states of the system P(S S) is the transition relation of the system the set of states may be infinite Xavier Rival Operational Semantics 4 / 50

Transition systems and small step semantics Definition and properties Transition systems: properties of the transition relation A deterministic system is such that a state fully determines the next state s 0, s 1, s 1 S, (s 0 s 1 s 0 s 1 ) = s 1 = s 1 Otherwise, a transition system is non deterministic, i.e.: s 0, s 1, s 1 S, s 0 s 1 s 0 s 1 s 1 s 1 Notes: transition relation defines atomic execution steps; it is often called small-step semantics or structured operational semantics steps are discrete (not continuous) to describe both discrete and continuous behaviors, we would need to look at hybrid systems (beyond the scope of this lecture) Xavier Rival Operational Semantics 5 / 50

Transition systems and small step semantics Definition and properties Transition systems: special states Initial / final states: we often consider transition systems with a set of initial and final states: a set of initial states S I S denotes states where the execution should start a set of final states S F S denotes states where the execution should reach the end of the program When needed, we add these to the definition of the transition systems ((S,, S I, S F ). Blocking state (not the same as final state): a state s 0 S is blocking when it is the origin of no transition: s 1 S, (s 0 s 1 ) example: we often introduce an error state (usually noted Ω to denote the erroneous, blocking configuration) Xavier Rival Operational Semantics 6 / 50

Transition systems and small step semantics Examples Outline 1 Transition systems and small step semantics Definition and properties Examples 2 Traces semantics 3 Summary Xavier Rival Operational Semantics 7 / 50

Transition systems and small step semantics Examples Finite automata as transition systems We can clearly formalize the word recognition by a finite automaton using a transition system: we consider automaton A = (Q, q i, q f, ) a state is defined by: the remaining of the word to recognize the automaton state that has been reached so far thus, S = Q L the transition relation of the transition system is defined by: (q 0, aw) (q 1, w) q 0 a q 1 the initial and final states are defined by: S I = {(q i, w) w L } S F = {(q f, ɛ)} Xavier Rival Operational Semantics 8 / 50

Transition systems and small step semantics Examples Pure λ-calculus A bare bones model of functional programing: λ-terms The set of λ-terms is defined by: t, u,... ::= x variable λx t abstraction t u application β-reduction (λx t) u β t[x u] if u β v then λx u β λx v if u β v then u t β v t if u β v then t u β t v The λ-calculus defines a transition system: S is the set of λ-terms and β the transition relation β is non-deterministic; example? though, ML fixes an execution order given a lambda term t 0, we may consider (S, β, S I ) where S I = {t 0 } blocking states are terms with no redex (λx u) v Xavier Rival Operational Semantics 9 / 50

Transition systems and small step semantics Examples A MIPS like assembly language: syntax We now consider a (very simplified) assembly language machine integers: sequences of 32-bits (set: B 32 ) instructions are encoded over 32-bits (set: I MIPS ) and stored into the same space as data (i.e., I MIPS B 32 ) Memory configurations program counter pc current instruction general purpose registers r 0... r 31 main memory (RAM) mem : Addrs B 32 where Addrs B 32 Instructions i ::= ( I MIPS ) add r d, r s0, r s1 addition addi r d, r s0, v add. v B 32 sub r d, r s0, r s1 subtraction b dst branch blt r s0, r s1, dst cond. branch ld r d, o, r x relative load st r d, o, r x relative store v, dst, o B 32 Xavier Rival Operational Semantics 10 / 50

Transition systems and small step semantics Examples A MIPS like assembly language: states Definition: state A state is a tuple (pc, ρ, µ) which comprises: a program counter value pc B 32 a function mapping each general purpose register to its value ρ : {0,..., 31} B 32 a function mapping each memory cell to its value µ : Addrs B 32 What would a dangerous state be? writing over an instruction reading or writing outside the program s memory we cannot fully formalize these yet... as we need to formalize the behavior of each instruction first Xavier Rival Operational Semantics 11 / 50

Transition systems and small step semantics Examples A MIPS like assembly language: transition relation We assume a state s = (pc, ρ, µ) and that µ(pc) = i; then: if i = add r d, r s0, r s1, then: s (pc + 4, ρ[d ρ(s0) + ρ(s1)], µ) if i = addi r d, r s0, v, then: s (pc + 4, ρ[d ρ(s0) + v], µ) if i = sub r d, r s0, r s1, then: s (pc + 4, ρ[d ρ(s0) ρ(s1)], µ) if i = b dst, then: s (dst, ρ, µ) Xavier Rival Operational Semantics 12 / 50

Transition systems and small step semantics Examples A MIPS like assembly language: transition relation We assume a state s = (pc, ρ, µ) and that µ(pc) = i; then: if i = blt r s0, r s1, dst, then: { (dst, ρ, µ) if ρ(s0) < ρ(s1) s (pc + 4, ρ, µ) otherwise if i = ld r d, o, r x, then: { (pc + 4, ρ[d µ(ρ(x) + o)], µ) if µ(ρ(x) + o) is defined s Ω otherwise if i = st r d, o, r x, then: { (pc + 4, ρ, µ[ρ(x) + o ρ(d)]) if µ(ρ(x) + o) is defined s Ω otherwise Xavier Rival Operational Semantics 13 / 50

Transition systems and small step semantics Examples A simple imperative language: syntax We now look at a more classical imperative language (intuitively, a bare-bone subset of C): Syntax variables X: finite, predefined set of variables labels L: before and after each statement values V: V int V float... e ::= v V int V float... e + e e e... expressions c ::= TRUE FALSE e < e e = e conditions i ::= x := e; assignment if(c) b else b condition while(c) b loop b ::= {i;... ; i; } block, program(p) Xavier Rival Operational Semantics 14 / 50

Transition systems and small step semantics Examples A simple imperative language: states A non-error state should fully describe the configuration at one instant of the program execution: the memory state defines the current contents of the memory m M = X V the control state defines where the program currently is analoguous to the program counter can be defined by adding labels L = {l0, l 1,...} between each pair of consecutive statements; then: S = L M {Ω} or by the program remaining to be executed; then: S = P M {Ω} Xavier Rival Operational Semantics 15 / 50

Transition systems and small step semantics Examples A simple imperative language: semantics of expressions The semantics e of expression e should evaluate each expression into a value, given a memory state Evaluation errors may occur: division by zero... error value is also noted Ω Thus: e : M V {Ω} Definition, by induction over the syntax: v (m) = v x (m) = m(x) e 0 + e 1 (m) = e { 0 (m)+ e 1 (m) Ω if e1 (m) = 0 e 0 /e 1 (m) = e 0 (m)/ e 1 (m) otherwise where is the machine implementation of operator, and is Ω-strict, i.e., v V, v Ω = Ω v = Ω. Xavier Rival Operational Semantics 16 / 50

Transition systems and small step semantics Examples A simple imperative language: semantics of conditions The semantics c of condition c should return a boolean value It follows a similar definition to that of the semantics of expressions: c : M V bool {Ω} Definition, by induction over the syntax: TRUE (m) = TRUE FALSE (m) = FALSE TRUE e 0 < e 1 (m) = FALSE Ω TRUE e 0 = e 1 (m) = FALSE Ω if e 0 (m) < e 1 (m) if e 0 (m) e 1 (m) if e 0 (m) = Ω or e 1 (m) = Ω if e 0 (m) = e 1 (m) if e 0 (m) e 1 (m) if e 0 (m) = Ω or e 1 (m) = Ω Xavier Rival Operational Semantics 17 / 50

Transition systems and small step semantics Examples A simple imperative language: transitions We now consider the transition induced by each statement. Case of assignment l 0 : x = e; l 1 if e (m) Ω, then (l 0, m) (l 1, m[x e (m)]) if e (m) = Ω, then (l 0, m) Ω Case of condition l 0 : if(c){l 1 : b t l 2 } else{l 3 : b f l 4 } l 5 if c (m) = TRUE, then (l 0, m) (l 1, m) if c (m) = FALSE, then (l 0, m) (l 3, m) if c (m) = Ω, then (l 2, m) (l 5, m) (l 4, m) (l 5, m) (l 0, m) Ω Xavier Rival Operational Semantics 18 / 50

Transition systems and small step semantics Examples A simple imperative language: transitions Case of loop l 0 : while(c){l 1 : b t l 2 } l 3 { (l0, m) (l if c (m) = TRUE, then 1, m) (l 2, m) (l 1, m) { (l0, m) (l if c (m) = FALSE, then 3, m) (l 2, m) (l 3, m) { (l0, m) Ω if c (m) = Ω, then (l 2, m) Ω Case of {l 0 : i 0 ; l 1 :... ; l n 1 i n 1 ; l n } the transition relation is defined by the individual instructions Xavier Rival Operational Semantics 19 / 50

Transition systems and small step semantics Examples Extending the language with non-determinism The language we have considered so far is a bit limited: it is deterministic: at most one transition possible from any state it does not support the input of values Changes if we model non deterministic inputs...... with an input instruction: i ::=... x := input() l 0 : x := input(); l 1 generates transitions v V, (l 0, m) (l 1, m[x v]) one instruction induces non determinism... with a random function: e ::=... rand() expressions have a non-deterministic semantics: e : M P(V {Ω}) rand() (m) = V v (m) = {v} c : M P(V bool {Ω}) all instructions induce non determinism Xavier Rival Operational Semantics 20 / 50

Transition systems and small step semantics Examples Semantics of real world programming languages C language: several norms: ANSI C 99, ANSI C 11, K&R... not fully specified: undefined behavior implementation dependent behavior: architecture (ABI) or implementation (compiler...) unspecified parts: leave room for implementation of compilers and optimizations formalizations in HOL (C 99), in Coq (CompCert C compiler) OCaml language: more formal...... but still with some unspecified parts, e.g., execution order Xavier Rival Operational Semantics 21 / 50

Definitions Outline 1 Transition systems and small step semantics 2 Traces semantics Definitions Finite traces semantics Fixpoint definition Compositionality Infinite traces semantics 3 Summary Xavier Rival Operational Semantics 22 / 50

Definitions Execution traces So far, we considered only states and atomic transitions We now consider program executions as a whole Definition: traces A finite trace is a finite sequence of states s 0,..., s n, noted s 0,..., s n An infinite trace is an infinite sequence of states s 0,... Besides, we write: S for the set of finite traces S ω for the set of infinite traces S = S S ω for the set of finite or infinite traces Xavier Rival Operational Semantics 23 / 50

Definitions Operations on traces: concatenation Definition: concatenation The concatenation operator is defined by: s 0,..., s n s 0,..., s n = s 0,..., s n, s 0,..., s n s 0,..., s n s 0,... = s 0,..., s n, s 0,... s 0,..., s n,... σ = s 0,..., s n,... We also define: the empty trace ɛ, neutral element for the length operator. : ɛ = 0 s 0,..., s n = n + 1 s 0,... = ω Xavier Rival Operational Semantics 24 / 50

Definitions Comparing traces: the prefix order relation Definition: prefix order relation Relation is defined by: s 0,..., s n s 0,..., s n { n n i 0, n, s i = s i s 0,... s 0,... i N, s i = s i s 0,..., s n s 0,... i 0, n, s i = s i Proof: straightforward application of the definition of order relations Xavier Rival Operational Semantics 25 / 50

Finite traces semantics Outline 1 Transition systems and small step semantics 2 Traces semantics Definitions Finite traces semantics Fixpoint definition Compositionality Infinite traces semantics 3 Summary Xavier Rival Operational Semantics 26 / 50

Finite traces semantics Semantics of finite traces We consider a transition system S = (S, ) Definition The finite traces semantics S is defined by: Example: S = { s 0,..., s n S i, s i s i+1 } contrived transition system S = ({a, b, c, d}, {(a, b), (b, a), (b, c)}) finite traces semantics: S = { ɛ, a, b,..., a, b, a, b, a,..., a, b, a, a, b,..., a, b, a, b, b, a,..., a, b, a, b, a, b,..., a, b, a, b, c, b, a,..., a, b, a, b, c c, d } Xavier Rival Operational Semantics 27 / 50

Finite traces semantics Interesting subsets of the finite trace semantics We consider a transition system S = (S,, S I, S F ) the initial traces, i.e., starting from an initial state: { s 0,..., s n S s 0 S I } the traces reaching a blocking state: {σ S σ S, σ σ = σ = σ } the traces ending in a final state: { s 0,..., s n S s n S F } Example (same transition system, with S I = {a} and S F = {c}): traces from an initial state ending in a final state: { a, b,..., a, b, a, b, c } Xavier Rival Operational Semantics 28 / 50

Finite traces semantics Example: finite automaton We consider the example of the previous course: L = {a, b} Q = {q 0, q 1, q 2 } a q i = q 0 q f = q 2 q a 0 q 1 q 2 a b a q 1 q 1 q 2 q 2 q 1 q 0 Then, we have the following traces: τ 0 = (q 0, ab), (q 1, b), (q 2, ɛ) τ 1 = (q 0, abab), (q 1, bab), (q 2, ab), (q 1, b), (q 2, ɛ) τ 2 = (q 0, ababab), (q 1, babab), (q 2, abab), (q 1, bab) τ 3 = (q 0, abaaa), (q 1, baaa), (q 2, aaa), (q 1, aa) Then: τ 0, τ 1 are initial traces, reaching a final state τ 2 is an initial trace, and is not maximal τ 3 reaches a blocking state, but not a final state Xavier Rival Operational Semantics 29 / 50 b

Finite traces semantics Example: λ-term We consider λ-term λy ((λx y)((λx x x) (λx x x))), and show two traces generated from it (at each step the reduced lambda is shown in red): Then: τ 0 = λy ((λx y)((λx x x) (λx x x))) λy y τ 1 = λy ((λx y)((λx x x) (λx x x))) λy ((λx y)((λx x x) (λx x x))) λy ((λx y)((λx x x) (λx x x))) τ 0 is a maximal trace; it reaches a blocking state (no more reduction can be done) τ 1 can be extended for arbitrarily many steps ; the second part of the course will study infinite traces Xavier Rival Operational Semantics 30 / 50

Finite traces semantics Example: imperative program Similarly, we can write the traces of a simple imperative program: l 0 : x := 1; l 1 : y := 0; l 2 : while(x < 4){ l 3 : y := y + x; l 4 : x := x + 1; l 5 : } l 6 : (final program point) τ = (l 0, x = 6, y = 8 ), (l 1, x = 1, y = 8 ), (l 2, x = 1, y = 0 ), (l 3, x = 1, y = 0 ), (l 4, x = 1, y = 1 ), (l 5, x = 2, y = 1 ), (l 3, x = 2, y = 1 ), (l 4, x = 2, y = 3 ), (l 5, x = 3, y = 3 ), (l 3, x = 3, y = 3 ), (l 4, x = 3, y = 6 ), (l 5, x = 4, y = 6 ), (l 6, x = 4, y = 6 ) very precise description of what the program does...... but quite cumbersome Xavier Rival Operational Semantics 31 / 50

Fixpoint definition Outline 1 Transition systems and small step semantics 2 Traces semantics Definitions Finite traces semantics Fixpoint definition Compositionality Infinite traces semantics 3 Summary Xavier Rival Operational Semantics 32 / 50

Fixpoint definition Towards a fixpoint definition We consider again our contrived transition system Traces by length: S = ({a, b, c, d}, {(a, b), (b, a), (b, c)}) i traces of length i 0 ɛ 1 a, b, c, d 2 a, b, b, a, b, c 3 a, b, a, b, a, b, a, b, c 4 a, b, a, b, b, a, b, a, b, a, b, c Like the automaton in lecture 1, this suggests a least fixpoint definition: traces of length i + 1 can be derived from the traces of length i, by adding a transition Xavier Rival Operational Semantics 33 / 50

Fixpoint definition Trace semantics fixpoint form We define a semantic function, that computes the traces of length i + 1 from the traces of length i (where i 1): Finite traces semantics as a fixpoint Let I = {ɛ} { s s S}. Let F be the function defined by: F : P(S ) P(S ) X X { s 0,..., s n, s n+1 s 0,..., s n X s n s n+1 } Then, F is continuous and thus has a least-fixpoint greater than I; moreover: lfp I F = S = n N F n (I) Xavier Rival Operational Semantics 34 / 50

Fixpoint definition Fixpoint definition: proof (1), fixpoint existence First, we prove that F is continuous. Let X P(S ) and A = X X X. Then: F ( X X X ) = A { s 0,..., s n, s n+1 ( s 0,..., s n X X X ) s n s n+1 } = A { s 0,..., s n, s n+1 ( X X, s 0,..., s n X ) s n s n+1 } = A { s 0,..., s n, s n+1 X X, s 0,..., s n X s n s n+1 } = ( X X X ) ( X X { s 0,..., s n, s n+1 s 0,..., s n X s n s n+1 } ) = X X (X { s 0,..., s n, s n+1 s 0,..., s n X s n s n+1 }) = X X F (X ) Function F is -complete, hence continuous. As (P(S ), ) is a CPO, the continuity of F entails the existence of a least-fixpoint (Kleene theorem); moreover, it implies that: lfp I F = n N F n (I) Xavier Rival Operational Semantics 35 / 50

Fixpoint definition Fixpoint definition: proof (2), fixpoint equality We now show that S is equal to lfp I F, by showing the property below, by induction over n: k n, s 0,..., s k F n (I) s 0,..., s k S at rank 0, only traces of length 1 need be considered: s S s S s F 0 (I) at rank n + 1, and assuming the property holds at rank n (the equivalence is obvious for traces of length 1): s 0,..., s k, s k+1 S s 0,..., s k S s k s k+1 s 0,..., s k F n (I) s k s k+1 (k n since k + 1 n + 1) s 0,..., s k, s k+1 F n+1 (I) Xavier Rival Operational Semantics 36 / 50

Fixpoint definition Trace semantics fixpoint form: example Example, with the same simple transition system S = (S, ): S = {a, b, c, d} is defined by a b, b a and b c Then, the first iterates are: F 0 (I) = {ɛ, a, b, c, d } F 1 (I) = F 0 (I) { b, a, a, b, b, c } F 2 (I) = F 1 (I) { a, b, a, b, a, b, a, b, c } F 3 (I) = F 2 (I) { b, a, b, a, a, b, a, b, b, a, b, c } F 4 (I) = F 3 (I) { a, b, a, b, a, b, a, b, a, b, a, b, a, b, c } F 5 (I) =... the traces of S of length n + 1 appear in F n (I) Xavier Rival Operational Semantics 37 / 50

Compositionality Outline 1 Transition systems and small step semantics 2 Traces semantics Definitions Finite traces semantics Fixpoint definition Compositionality Infinite traces semantics 3 Summary Xavier Rival Operational Semantics 38 / 50

Compositionality Notion of compositional semantics The traces semantics definition we have seen is global: the whole system defines a transition relation we iterate this relation until we get a fixpoint Though, a modular definition would be nicer, to allow reasoning on program fragments, or derive properties of a program from properties of its pieces... Can we derive a more modular expression of the semantics? Xavier Rival Operational Semantics 39 / 50

Compositionality Notion of compositional semantics Observation: programs often have an inductive structure λ-terms are defined by induction over the syntax imperative programs are defined by induction over the syntax there are exceptions: our MIPS language does not naturally look that way Definition: compositional semantics A semantics. is said to be compositional when the semantics of a program can be defined as a function of the semantics of its parts, i.e., When program π writes down as C[π 0,..., π k ] where π 0,..., π k are its components, there exists a function F C such that π = F C ( π 0,..., π k ), where F C depends only on syntactic construction F C. Xavier Rival Operational Semantics 40 / 50

Compositionality Case of a simplified imperative language Case of a sequence of two instructions b l 0 : i 0 ; l 1 : i 1 ; l 2 : b = i 0 i 1 { s 0,..., s m n 0, m, s 0,..., s n i 0 s n,..., s m i 1 } This amounts to concatenating traces of i 0 and i 1 that share a state in common (necessarily at point l 1 ). Cases of a condition, a loop: similar by concatenation of traces around junction points by doing a least-fixpoint computation over loops We can provide a compositional semantics for our simplified imperative language Xavier Rival Operational Semantics 41 / 50

Compositionality Case of λ-calculus Case of a λ-term t = (λx u) v: executions may start with a reduction in u executions may start with a reduction in v executions may start with the reduction of the head redex an execution may mix reductions steps in u and v in an arbitrary order No nice compositional trace semantics of λ-calculus... Xavier Rival Operational Semantics 42 / 50

Infinite traces semantics Outline 1 Transition systems and small step semantics 2 Traces semantics Definitions Finite traces semantics Fixpoint definition Compositionality Infinite traces semantics 3 Summary Xavier Rival Operational Semantics 43 / 50

Infinite traces semantics Non termination Can the finite traces semantics express non termination? Consider the case of our contrived system: S = {a, b, c, d} ( ) = {(a, b), (b, a), (b, c)} this system clearly has non-terminating behaviors: it can loop from a to b and back forever the finite traces semantics does show the existence of this cycle as there exists an infinite chain of finite traces for the prefix order : a, b, a, b, a, a, b, a, b, a, b, a, b, a,... S though, the existence of this chain is not very obvious Thus, we now define a semantics made of infinite traces Xavier Rival Operational Semantics 44 / 50

Infinite traces semantics Semantics of infinite traces We consider a transition system S = (S, ) Definition The infinite traces semantics S ω is defined by: S ω = { s 0,... S ω i, s i s i+1 } Infinite traces starting from an initial state (considering S = (S,, S I, S F )): { s 0,... S ω s 0 S I } Example: contrived transition system defined by S = {a, b, c, d} ( ) = {(a, b), (b, a), (b, c)} the infinite traces semantics contains exactly two traces S ω = { a, b,..., a, b, a, b,..., b, a,..., b, a, b, a,... } Xavier Rival Operational Semantics 45 / 50

Infinite traces semantics Fixpoint form Can we also provide a fixpoint form for S ω? Intuitively, s 0, s 1,... S ω if and only if n, s n s n+1, i.e., Let F ω be defined by: n N, k n, s k s k+1 F ω : P(S ω ) P(S ω ) X { s 0, s 1,..., s n,... s 1,..., s n,... X s 0 s 1 } Then, we can show by induction that: σ S ω n N, σ F n ω(s ω ) n N F n ω(s ω ) Xavier Rival Operational Semantics 46 / 50

Infinite traces semantics Fixpoint form of the semantics of infinite traces Infinite traces semantics as a fixpoint Let F ω be the function defined by: F ω : P(S ω ) P(S ω ) X { s 0, s 1,..., s n,... s 1,..., s n,... X s 0 s 1 } Then, F ω is -continuous and thus has a greatest-fixpoint; moreover: gfp S ωf ω = S ω = n N F n ω(s ω ) Proof sketch: the -contiunity proof is similar as for the -continuity of F by the dual version of Kleene s theorem, gfp S ωf ω exists and is equal to n N F n ω(s ω ), i.e. to S ω (similar induction proof) Xavier Rival Operational Semantics 47 / 50

Infinite traces semantics Fixpoint form of the infinite traces semantics: iterates Example, with the same simple transition system: S = {a, b, c, d} is defined by a b, b a and b c Then, the first iterates are: Fω(S 0 ω ) = S ω Fω(S 1 ω ) = a, b S ω b, a S ω b, c S ω Fω(S 2 ω ) = b, a, b S ω a, b, a S ω a, b, c S ω Fω(S 3 ω ) = a, b, a, b S ω b, a, b, a S ω b, a, b, c S ω Fω(S 4 ω ) =... Intuition at iterate n, prefixes of length n + 1 match the traces in the infinite semantics only a, b,..., a, b, a, b,... and b, a,..., b, a, b, a,... belong to all iterates Xavier Rival Operational Semantics 48 / 50

Summary Outline 1 Transition systems and small step semantics 2 Traces semantics 3 Summary Xavier Rival Operational Semantics 49 / 50

Summary Summary We have discussed: small-step / structural operational semantics: individual program steps big-step / natural semantics: program executions as sequences of transitions their fixpoint definitions and properties Next lectures: another family of semantics, more compact and compositional semantic program and proof methods Xavier Rival Operational Semantics 50 / 50