An That Proves Markov’s Principle (revision fixing typos – April 2021)

Hugo Herbelin INRIA - PPS Paris, France [email protected]

Abstract—We design an intuitionistic predicate logic that was shown e.g. by Smorynski in [2]. One of our objective supports a limited amount of classical reasoning, just enough is then to study intuitionistic logic extended with Markov’s to prove a variant of Markov’s principle suited for predicate principle as a proper constructive logic. logic. At the computational level, the extraction of an existential An elementary method to prove the intuitionistic admis- witness out of a proof of its double negation is done by using sibility of Markov’s rule has been given by Friedman who a form of statically-bound exception mechanism, what can be introduced for this purpose a notion of A-translation [9] (see seen as a direct-style variant of Friedman’s A-translation. also Dragalin [10]). In fact, Friedman’s proof does not prove Keywords-Markov’s principle; intuitionistic logic; proof-as- ¬¬∃x A(x) → ∃x A(x) for any propositional formula A(x) program correspondence; exceptions but only for those A(x) that do not contain the implication connective (this was noticed for instance in Berger [11]). I.INTRODUCTION More generally, Friedman’s method can be interpreted from In arithmetic, Markov’s principle is a weak classical the proof-theoretical point of view as a proof that ¬¬A → A scheme stating ¬¬∃x A(x) → ∃x A(x) for every decidable is admissible in intuitionistic predicate logic whenever A is formula A(x). Though not derivable [1] in Heyting Arith- a ∀-→-free formula, the case of ∃x A(x) for A(x) →-free metic (i.e. intuitionistic arithmetic), its formulation as a rule being just a prototypical instance of the scheme. (Markov’s rule) is admissible (see e.g. [2] for a classical We concentrate here on predicate logic and refer to in proof). One of the important properties of Markov’s princi- this context as Markov’s principle the scheme that asserts ple is that it implies, via the double-negation translation, the ¬¬T → T whenever T is ∀-→-free. Extending intuitionistic conservativity of classical logic over intuitionistic logic for logic with classical reasoning on ∀-→-free formulas is then 0 Π2-formulas. the same as adding Markov’s principle and we show that the Markov’s principle is realizable (in the sense of resulting logic is constructive in the sense that it satisfies Kleene [3]) by unbounded search (ensured terminating by normalization and the disjunction and existence properties. use of classical reasoning). It is also realizable along Gödel At the computational level, the added classical rules can “Dialectica” functional interpretation [4] with witness ob- be seen as a mechanism of (statically-bound) exception tained by selecting the first suitable instance among those throwing similar to the one introduced by Nakano in his occurring in the proof of ¬¬∃x A(x). This is generally by (intuitionistic) catch and throw calculus [12]1. one of these two ways that programs are extracted from II. IQC :AN INTUITIONISTIC PREDICATE LOGIC proofs that use Markov’s principle, especially in constructive MP THAT PROVES MARKOV’S PRINCIPLE analysis where Markov’s principle amounts to the property ¬¬x 6= y → x 6= y (6= asserts the existence of a rational Usual intuitionistic predicate logic (IQC) is defined from number that separates x and y, see e.g. [5]). a set of function symbols f, g, . . . , each of a given arity and Howard’s correspondence [6] between natural deduction a set of predicate symbols P , Q, R, ... each with an arity and λ-calculus emphasized that Prawitz’s normalization of too. Functions symbols of arity 0 are called constants and intuitionistic natural deduction [7] can be used as a method predicate symbols of arity 0 are called atomic propositions. to compute with proofs. By later showing a connection Terms are built from a set of variables x, y, ... by between classical axioms and control operators, Griffin [8] t, u ::= f(~t) | x revealed that normalization is a computation device also for classical logic. While intuitionistic logic satisfies the where f ranges over function symbols and ~t denotes in f(~t) disjunction and existence properties, it is however not the a sequence of terms of length the arity of f. Formulas are case of classical logic and Markov’s principle has this 1Ordinary (dynamically bound) exceptions, as found in the ML, C++ or peculiarity to lie in between, being classical in its form and Java programming languages can actually be considered too, this is part of still retaining the disjunction and existence properties, as ongoing work. A ∈ Γ AXIOM Γ `∆ A

Γ `∆ A1 Γ `∆ A2 Γ `∆ A1 ∧ A2 i ∧I ∧E Γ `∆ A1 ∧ A2 Γ `∆ Ai

Γ `∆ Ai Γ `∆ A1 ∨ A2 Γ,A1 `∆ B Γ,A2 `∆ B i ∨I ∨E Γ `∆ A1 ∨ A2 Γ `∆ B

Γ,A `∆ B Γ `∆ A → B Γ `∆ A →I →E Γ `∆ A → B Γ `∆ B

Γ `∆ A(x) x fresh Γ `∆ ∀x A(x) ∀I ∀E Γ `∆ ∀x A(x) Γ `∆ A(t)

Γ `∆ A(t) Γ `∆ ∃x A(x)Γ,A(x) `∆ B x fresh ∃I ∃E Γ `∆ ∃x A(x) Γ `∆ B

Γ `∆ ⊥ >I ⊥E Γ `∆ > Γ `∆ C

Γ `T,∆ T Γ `∆ TT ∈ ∆ CATCH THROW Γ `∆ T Γ `∆ C

Figure 1. Inference rules of IQCMP

built from the standard connectives and quantifiers by the The main difference between IQCMP and IQC is that grammar the former supports classical reasoning on ∀-→-free formu- las. This is implemented by the rules CATCH and THROW A, B ::= P (~t) | > | ⊥ | A → B | A ∧ B | A ∨ B which say that to prove a ∀-→-free formula T , one is | ∀x A | ∃x A allowed to change its mind during the proof and to restart a where P ranges over predicate symbols and ~t is a sequence new proof of T at any time. of terms whose length is the arity of P . Negation ¬A is The main properties of IQCMP is that it proves Markov’s defined as A → ⊥. In ∀x A and ∃x A, x is bound and principle while still retaining the disjunction and existence freely subject to renaming (so-called α-conversion). properties that are characteristic of intuitionistic logic. IQCMP is an extension of IQC. Its inference rules are Let us write MP for the scheme ¬¬T → T where T given on Figure 1 (we use natural deduction). The subclass is a ∀-→-free formula (this is the “proof-theoretic” way to of ∀-→-free formulas plays a special role and we use T , U, think about Markov’s principle) and MP∃ for the scheme ... to denote such formulas: ¬¬∃~xA(~x) → ∃~xA(~x) where A(~x) is an →-free propo- sition (this is the “standard” way to think about Markov’s T,U ::= P (~t) | > | ⊥ | T ∧ U | T ∨ U | ∃x T principle). Both schemes are logically equivalent. Contexts of formulas, written Γ, are ordered sequences of Theorem 1: In IQCMP , MP and MP∃ hold. formulas. Contexts of ∀-→-free formulas, written ∆, are Proof: One gets a proof of T `T ⊥ by applying ordered sequences of ∀-→-free formulas. By ¬∆ is meant THROW. By →I and →E we obtain a proof of ¬¬T `T ⊥. the context obtained by distributing ¬ over the formulas of By applying ⊥E followed by CATCH, we get a proof of ∆. Note that IQC can be characterized as the subset of ¬¬T ` T from which ¬¬T → T derives (note that we 0 IQCMP obtained by removing the rules CATCH and THROW freely use the lemma that Γ `∆ A implies Γ `∆0 A for and keeping ∆ empty. Γ ⊂ Γ0 and ∆ ⊂ ∆0; this weakening lemma, as expected, (a : A) ∈ Γ AXIOM Γ `∆ a : A

Γ `∆ p1 : A1 Γ `∆ p2 : A2 Γ `∆ p : A1 ∧ A2 i ∧I ∧E Γ `∆ (p1, p2): A1 ∧ A2 Γ `∆ πip : Ai

Γ `∆ p : Ai Γ `∆ p : A1 ∨ A2 Γ, a1 : A1 `∆ p1 : B Γ, a2 : A2 `∆ p2 : B i ∨I ∨E Γ `∆ ιi(p): A1 ∨ A2 Γ `∆ case p of [a1.p1 | a2.p2]: B

Γ, a : A `∆ p : B Γ `∆ p : A → B Γ `∆ q : A →I →E Γ `∆ λa.p : A → B Γ `∆ p q : B

Γ `∆ p : A(x) x fresh Γ `∆ p : ∀x A(x) ∀I ∀E Γ `∆ λx.p : ∀x A(x) Γ `∆ pt : A(t)

Γ `∆ p : A(t) Γ `∆ p : ∃x A(x)Γ, a : A(x) `∆ q : B x fresh ∃I ∃E Γ `∆ (t, p): ∃x A(x) Γ `∆ dest p as (x, a) in q : B

Γ `∆ p : ⊥ >I ⊥E Γ `∆ () : > Γ `∆ efq p : C

Γ `α:T,∆ p : T Γ `∆ p : T (α : T ) ∈ ∆ CATCH THROW Γ `∆ catchα p : T Γ `∆ throwα p : C

Figure 2. Proof-term annotation of IQCMP

indeed holds in IQCMP ). The case of MP∃ holds by logical The language of proofs is defined by the grammar equivalence with MP . p, q ::= a | ι (p) | (p, q) | (t, p) | λa.p | λx.p | () Theorem 2: Γ ` A in IQC iff MP, Γ ` A in IQC i MP | p [a .p | a .p ] iff MP , Γ ` A in IQC. case of 1 1 2 2 ∃ | π (p) | dest p as (x, a) in q Proof: By Theorem 1, Γ,MP ` A implies Γ ` A i | p q | p t | efq p which obviously is a proof of IQCMP . Conversely, we | catchαp | throwαp prove by induction on a derivation of Γ `∆ A in IQCMP that MP, Γ, ¬∆ ` A holds in IQC. All cases are direct and where a, b, . . . range over a first set of proof variables. we use MP for interpreting the rule CATCH. and α, β, . . . range over another set of proof variables. Theorem 3 (Disjunction property): In IQCMP , if The constructions λa.p, case p of [a1.p1 | a2.p2] and ` A1 ∨ A2 then ` A1 or ` A2. dest p as (x, a) in q bind a, a1 and a2. The constructions Theorem 4 (Existence property): In IQCMP , if λx.p and dest p as (x, a) in q bind x. The construction ` ∃x A(x) then there exists t such that ` A(t). catchαp binds α. The binders are considered up to the The proof of these last two theorems is the subject of the actual name used to represent the binder (α-conversion). next section. The annotation of IQCMP with proof-terms is given in Figure 2 where the contexts Γ and ∆ are now maps III.THE PROOF THEORY OF IQCMP from variable names to formulas. For instance, the proof of Markov’s principle in Theorem 1 is We show that the proofs of IQCMP have a computational interpretation as programs in a λ-calculus extended with a λa.catchα efq (a λb.throwα b) mechanism of statically-bound exceptions implemented with operators named catch and throw. with the typing derivation given on Figure 3. AXIOM b : T `α:T b : T THROW b : T `α:T throwα b : ⊥ AXIOM →I a : ¬¬T `α:T a : ¬¬T `α:T λb.throwα b : ¬T →E a : ¬¬T `α:T a (λb.throwα b): ⊥ ⊥E a : ¬¬T `α:T efq a (λb.throwα b): T CATCH a : ¬¬T ` catchα efq a (λb.throwα b): T →I ` λa.catchα efq a (λb.throwα b): ¬¬T → T

Figure 3. Proof of MP

A subclass of proofs will play a particular role in extract- been proved using Markov’s principle in the empty context, ing the intuitionistic content of weakly classical proofs of its call-by-value evaluation reveals that the call to Markov’s IQCMP . These are the values defined by principle is in fact useless. The operators catch and throw behave like the similarly V ::= a | ι (V ) | (V,V ) | (t, V ) | λa.p | λx.p | () i named operators of Nakano [12] or Crolard [13]. Like Another class of expressions will be useful to define the in [12], but on the contrary of [13] (or of Parigot’s λµ- reduction, it is the class of elementary evaluation contexts calculus [14] to which the calculus of [13] is equivalent), defined by catch does not capture its surrounding context (i.e. there is no rule of the form F [catchαp] → catchβ p[throwα [] ← F [ ] ::= case [] of [a1.p1 | a2.p2] throwβ F [ ]]). As such, throw behaves as an exception | πi([ ]) | dest [] as (x, a) in p raiser and catch as an exception handler but still not as | [] q | (λx.q)[] in standard programming languages like Java or ML, since | [] t | efq [] | throwα [] there exceptions are dynamically bound (i.e. the substitution | ιi([ ]) | ([ ], p) | (V, [ ]) | (t, [ ]) is not capture-free) while in IQCMP they are statically- 2 For F [] an elementary evaluation context and p a proof, bound (i.e. the substitution is capture-free) . Alternatively, we write F [p] for the proof obtained by plugging p into the catchαp can be seen as a blocked control operator (i.e. hole of F []. as an expression of the form # callccα p where # is a delimiter [15] that blocks the interaction of callccα p with We can now define evaluation in IQCMP as the congruent closure of the following reductions: its outside and expects it first to evaluate – to a value – before being observed by the evaluation context of the delimiter). (λa.p) V → p[a ← V ] We now check that the reduction system is compatible (λx.p) t → p[x ← t] with typing. ι (V ) [a .p | a .p ] → p [a ← V ] case i of 1 1 2 2 i i Theorem 5 (Strengthening): If Γ ` V : T then (t, V ) (x, a) p → p[x ← t][a ← V ] ∆,T dest as in Γ ` V : T . π (V ,V ) → V ∆ i 1 2 i Proof: Obvious since the syntax of V refers to no p F [efq p] → efq p (and hence to no catch or throw) as soon as → and ∀ are F [throw p] → throw p α α excluded in T . catchαthrowαp → catchαp Theorem 6 (Subject reduction): If Γ `∆ p : A and p → q catchαthrowβV → throwβ V (α 6= β) then Γ `∆ q : A. catchαV → V Proof: By checking all cases, using Strengthening for where the substitutions p[a ← V ] and p[x ← t] are capture- the last two rules. Note that since catch, on the contrary of free with respect to the three kinds of variables (x, a and standard classical operators like callcc, does not capture its α). context, its type remains unchanged and the ∀-→ constraint Note that this is a call-by-value reduction semantics and on the formulas of ∆ is preserved. that we do not consider commutative cuts just because we We then characterize the set of normal forms in IQCMP . are only concerned with the normalization of closed proofs Theorem 7 (Characterization of normal forms): The set and commutative cuts are not needed for that purpose. of typed normal forms for → corresponds to the entry r Note also that because catchα only applies to proofs of ∀-→-free formulas, V cannot contain subterms of the form λa.p or λx.p in the last two rules and one is sure that α 2Compare substitution with capture (λa.catchα(a, throwα1))(throwα2) → catchα(throwα2, throwα1) does not occur in V . This is a crucial point of the design of to capture-free substitution the system which ensures that when a ∀-→-free formula has (λa.catchα(a, throwα1))(throwα2) → catchβ (throwα2, throwβ 1). of the following grammar: call-by-name as for call-by-value by embedding either into or into Urban’s non-deterministic LK [19]) or r ::= W | throw W | s | efq s+ | throw s α α LK [20] (for which a non-deterministic cut-elimination W ::= a | ι (W ) | (W, W ) | (t, W ) | λx.r | λa.r | () µµ˜ i able to capture call-by-value exists [21])3. s ::= π (s+) | s+ r | s+ t | (λa.r) s i We are now ready to prove Theorems 3 and 4. Given | ι (s) | (s, r) | (W, s) | (t, s) i a proof of ` p : A ∨ A , we know by progress and | case s+ of [a .r | a .r ] 1 2 1 1 2 2 normalization that p eventually reduces to a value V which, | dest s+ as (x, a) in r by subject reduction, satisfies ` V : A ∨ A . By inspection | catch s | catch efq s+ | catch throw s 1 2 α α α β of the possible forms of V , we know that we have either s+ ::= s | a a proof of ` A1 or a proof of ` A2. Similarly, from (we have to add efq W if we want to characterize untyped ` p : ∃x A(x) we derive ` V : ∃x A(x) for some V and normal forms.) hence ` A(t) for some term t. Proof: By inspection of the form of proofs that are not reducible. If the normal proof is not a normal value W nor IV. DISCUSSIONAND RELATION TO OTHER WORKS a throwαW waiting to erase its context up to its binding The codereliction of differential proof nets: In terms catchα, it is a proof whose computation is blocked by a of polarity in linear logic [23], the ∀-→-free constraint free variable. Such a normal proof either remains normal characterizes the formulas of intuitionistic logic that can be when placed into a normal context F [], in which case it built only from positive connectives (⊕, ⊗, 0, 1, !) and the has the form s, or it starts with throw or efq and captures why-not connective (“?”). In this framework, Markov’s prin- its surrounding context. A normal proof in s cannot be a ciple expresses that from such a ∀-→-free formula A (e.g. variable (otherwise it would be a value) but as soon as an ? ⊕x (?A(x)⊗?B(x))) where the presence of “?” indicates elimination rule is traversed, a variable can occur (entry s+). that the proof possibly used weakening (efq or throw) or contraction (catch), a linear proof of A purged from the We then check that the reduction system is not too occurrences of its “?” connective can be extracted (meaning rudimentary and that it at least produces head-normal form for the example above a proof of ⊕x(A(x) ⊗ B(x))). on closed proofs. Interestingly, the removal of the “?”, i.e. the steps from Theorem 8 (Progress): If ` p : A and p is not a (closed) ?P to P , correspond to applying the codereliction rule of value then p is reducible. differential proof nets [24]. Proof: According to Theorem 7, closed normal forms Nakano and Crolard’s intuitionistic catch and throw are necessarily in the set W . But this set is a subset of the calculi: Our calculus is very similar to the intuitionistic one set of values. proposed by Nakano [12]. However, in [12], the rule of intro- Theorem 9 (Normalization): If Γ `∆ p : A then p is duction of implication requires ∆ to be empty what prevents normalizable. from deriving Markov’s principle. Actually, as expressed Proof: We only give a hint. IQCMP is a subsystem of in Theorem 5 of [12], the logical expressiveness Nakano’s classical natural deduction equipped with a constrained call- calculus is the one of LJ. Another variant of intuitionistic by-value reduction system. Then, its normalization directly logic with control operators that does not increase the logical derives from the strong normalization of classical natural expressiveness can also be found in Crolard [25]. deduction, represented e.g. as call-by-value simply-typed Friedman’s A-translation: Expressed in our calculus, Parigot’s λµ-calculus [14] equipped with sums, products and Friedman’s A-translation [9] maps a proof of Γ `∆ A in the empty type, and with the appropriate reduction rules: IQCMP to a proof of Γ∆ ` A∆ in IQC, where B∆ is catchα p is interpreted by µα.[α]p, throwα p by µδ.[α]p for obtained by replacing each atom of B by the disjunction of δ fresh and efq p by µδ.[tp⊥]p where tp⊥ is the continuation the formulas in ∆. Through this translation, the THROW rule constant associated to the empty type (see [16]). It is unclear (assuming w.l.o.g that it is used with atomic conclusions) is whether normalization of this particular system has already interpretable as an injection. Also, for B ∀-→-free, we have W been considered in the literature. The closest work seems to B∆ → B ∨ ∆ from what we see that the CATCH rule is W be David and Nour [17] who have a normalization proof for interpretable (the proof of B∆ → B ∨ ∆ can be logically a symmetric λµ-calculus. This subsumes the call-by-value seen as the property that the proof of a ∀-→ formula case but only implication is considered. Alternatively, we can B can be purified from all its calls to throw by locally reduce normalization in natural deduction to normalization following a call-by-value reduction strategy; interestingly, in sequent calculus and embed the sequent calculus ver- when, B is a conjunction, there are two asymmetrical proofs sion of IQCMP into Danos-Joinet-Schellinx (call-by-value) 3Another way to prove the normalization of IQC without transiting LKQ [18] which (up to details about the representation MP by full classical logic is by embedding into IQC using a variant of of the axiom and contraction rules) is a subsystem of Coquand-Hofmann’s own variant [22] of Friedman’s translation, as done both LKtq [18] (which has cut-elimination as well for in a work in progress. W of C∆ ∧ D∆ → (C ∧ D) ∨ ∆ which match the two that Markov’s principle is necessary as soon as the con- asymmetrical ways to evaluate a pair along call-by-value nective ⊥ of the logic is interpreted as absurdity in the reduction). Henceforth, our calculus can be seen as a direct- model (a similar phenomenon happens for the completeness style variant4 of A-translation in the same way as callcc of intuitionistic logic for which the interpretation of ⊥ has provides with a direct-style representation of continuation- to be weakened so as to obtain intuitionistic proofs; see passing-style translation [8], [26]. e.g. Veldman [30]). We believe that both completeness for Independence of premise: The principle of indepen- classical logic and for intuitionistic logic could be carried dence of premise (IP) is the scheme (¬B → ∃x A(x)) → out in an extension of IQCMP to second-order arithmetic ∃x (¬B → A(x)). As a rule, the independence of premise without having to weaken the interpretation of ⊥. principle is admissible in intuitionistic logic [27] but it V. CONCLUSION is not admissible in IQCMP because if it were, taking A(x) an arbitrary atomic formula, one would obtain from We showed that adding classical reasoning on ∀-→-free MP that ∃x (¬¬∃y A(y) → A(x)), then, by intuitionistic formulas to intuitionistic logic preserves the intuitionistic reasoning, ∃x (∃y A(y) → A(x)), then, by Theorem 4, character of the logic, as witnessed by the preservation of ∃y A(y) → A(t) for some term t. But by Theorem 7, the disjunction and existence properties, while providing one sees that no normal proof can have this type (the with an effective intuitionistic proof of Markov’s principle. same reasoning holds in Heyting Arithmetic taking for A To compute with Markov’s principle, we used a form of a formula such that neither ` ¬¬A nor ` ¬A holds). statically-bound exception mechanism. Markov’s principle in arithmetic: Since any decidable ACKNOWLEDGMENTS formula can be expressed in terms of bounded existential quantification, conjunction and disjunction over decidable I thank Andreas Abel, Tristan Crolard, Danko Ilik, Guil- atoms, and hence as a ∀-→-free formula, we believe that laume Munch–Maccagnoni and Noam Zeilberger for fruitful by using an axiom-free presentation of Heyting Arithmetic, discussions on this topic. one could directly extend IQCMP to the arithmetic case. We would then get a constructive content of ¬¬∃x A(x) → REFERENCES ∃x A(x) for A(x) decidable which, in contrast to the re- [1] G. Kreisel, “The non-derivability of ¬ (x) A(x) → (Ex) ¬ alizer that successively checks the truth of each instance A(x), A primitive recursive, in intuitionistic formal systems of A(n), would not reconstruct a witness from scratch but (abstract),” The Journal of Symbolic Logic, vol. 23, no. 4, pp. 456–461, dec 1958. extract it from the proof. Especially, our constructivization of Markov’s principle not only contains its own proof of [2] A. S. Troelstra, Metamathematical Investigation of Intuition- termination but it also directly evaluates to the witness of istic Arithmetic and Analysis, ser. Lecture Notes in Mathe- the existential quantification. matics. Berlin: Springer-Verlag, 1973, vol. 344. Modified realizability: Markov’s principle is not real- izable with respect to Kreisel’s modified realizability [28] [3] S. C. Kleene, “On the interpretation of intuitionistic number theory,” The Journal of Symbolic Logic, vol. 10, no. 4, pp. which therefore does not suit to IQCMP . When interpreting 109–124, 1945. an implication, we have to anticipate the possibility that the function throw an exception. This suggests to introduce a [4] K. Gödel, “Über eine bisher noch nicht benützte Erweiterung modification of modified realizability where realizability is des finiten Standpunktes,” Dialectica, vol. 12, no. 3, pp. 280– parametrized by a set ∆ of possible ∀-→-free formulas a 287, Dec. 1958. program is allowed to escape to and to have that t realizes [5] U. Kohlenbach, Applied Proof Theory: Proof Interpretations A → B with respect to ∆ whenever t u realizes B with and their Use in Mathematics, ser. Monographs in Mathemat- respect to ∆ for u realizing A with respect to any extension ics. Springer, 2008. of ∆, using a kind of polymorphism similar to the one Coquand and Hofmann [22] introduced to extend Friedman’s [6] W. A. Howard, “The formulae-as-types notion of construc- tions,” in to H.B. Curry: Essays on Combinatory Logic, A-translation. Lambda Calculus and Formalism. Academic Press, 1980, Completeness proofs: Gödel and Kreisel proved that unpublished manuscript of 1969. completeness for classical predicate logic implies Markov’s principle. More precisely, Berardi and Valentini [29] showed [7] D. Prawitz, Natural Deduction, a Proof-Theoretical Study. Almqvist and Wiksell, Stockholm, 1965. 4Ongoing work suggests that Friedman’s A-translation is more pre- cisely related to a reduction system where exceptions are dynamically- [8] T. G. Griffin, “The formulae-as-types notion of control,” bound (instead of statically-bound) and where the underlying structure in Conf. Record 17th Annual ACM Symp. on Principles of (beside the call-by-value part used to reduce catch) is call-by-name. For Programming Languages, POPL ’90, San Francisco, CA, statically-bound exceptions, the relation has to be sought towards Coquand- USA, 17-19 Jan 1990. ACM Press, New York, 1990, pp. Hofmann’s variant [22] of Friedman’s A-translation. 47–57. [9] H. Friedman, “Classically and intuitionistically provably [19] C. Urban, “Classical logic and computation,” Ph.D. Thesis, recursive functions,” in Higher Set Theory, ser. Lecture University of Cambridge, Oct. 2000. Notes in Mathematics, D. S. Scott and G. H. Muller, Eds. Berlin/Heidelberg: Springer, 1978, vol. 669, pp. 21–27. [20] P.-L. Curien and H. Herbelin, “The duality of computation,” in Proceedings of the Fifth ACM SIGPLAN International [10] A. G. Dragalin, “New kinds of realizability and Markov’s Conference on Functional Programming, ICFP 2000, Mon- rule,” Soviet Mathematical Doklady, vol. 251, pp. 534–537, treal, Canada, September 18-21, 2000, ser. SIGPLAN Notices 1980. 35(9). ACM, 2000, pp. 233–243.

[11] U. Berger, “A computational interpretation of open induction,” [21] E. Polonovski, “Subsitutions explicites, logique et normalisa- in 19th IEEE Symposium on Logic in Computer Science (LICS tion,” Ph.D. Thesis, University Paris 7, Jun. 2003. 2004), 14-17 July 2004, Turku, Finland, Proceedings. IEEE Computer Society, 2004, p. 326. [22] T. Coquand and M. Hofmann, “A new method for establishing conservativity of classical systems over their intuitionistic ver- [12] H. Nakano, “A constructive formalization of the catch and sion,” Mathematical Structures in Computer Science, vol. 9, throw mechanism,” in Proceedings, Seventh Annual IEEE no. 4, pp. 323–333, 1999. Symposium on Logic in Computer Science, 22-25 June 1992, Santa Cruz, California, USA. IEEE Computer Society, 1992, [23] J.-Y. Girard, “Linear logic,” Theor. Comput. Sci., vol. 50, pp. pp. 82–89. 1–102, 1987.

[13] T. Crolard, “A confluent lambda-calculus with a catch/throw [24] T. Ehrhard and L. Regnier, “Differential interaction nets,” mechanism,” J. Funct. Program., vol. 9, no. 6, pp. 625–647, Electr. Notes Theor. Comput. Sci., vol. 123, pp. 35–74, 2005. 1999. [25] T. Crolard, “Extension de l’isomorphisme de Curry-Howard [14] M. Parigot, “Lambda-mu-calculus: An algorithmic interpre- au traitement des exceptions,” Ph.D. Thesis, University Paris tation of classical natural deduction,” in Logic Program- 7, Dec. 1996. ming and Automated Reasoning: International Conference LPAR ’92 Proceedings, St. Petersburg, Russia. Springer- [26] C. Murthy, “Extracting constructive content from classical Verlag, 1992, pp. 190–201. proofs,” Ph.D. Thesis, Cornell University, 1990.

[15] M. Felleisen, “The theory and practice of first-class prompts,” [27] A. Visser, “Aspects of diagonalization and provability,” Ph.D. in Proceedings of the 15th ACM Symposium on Principles of Thesis, University of Utrecht, Department of Philosophy, The Programming Languages (POPL ’88). ACM Press, New Nederland, Nov. 1981. York, Jan. 1988, pp. 180–190. [28] G. Kreisel, “Interpretation of analysis by means of functionals [16] Z. M. Ariola and H. Herbelin, “Minimal classical logic and of finite type,” in Constructivity in Mathematics, A. Heyting, control operators,” in Thirtieth International Colloquium on Ed. North-Holland, 1959, pp. 101–128. Automata, Languages and Programming , ICALP’03, Eind- hoven, The Netherlands, June 30 - July 4, 2003, ser. Lecture [29] S. Berardi and S. Valentini, “Krivine’s intuitionistic proof of Notes in Computer Science, vol. 2719. Springer-Verlag, classical completeness (for countable languages),” Ann. Pure 2003, pp. 871–885. Appl. Logic, vol. 129, no. 1-3, pp. 93–106, 2004.

[17] R. David and K. Nour, “Arithmetical proofs of strong nor- [30] W. Veldman, “An intuitionistic completeness theorem for malization results for symmetric λ-calculi,” Fundam. Inform., intuitionistic predicate logic,” J. Symb. Log., vol. 41, no. 1, vol. 77, no. 4, pp. 489–510, 2007. pp. 159–166, 1976.

[18] V. Danos, J.-B. Joinet, and H. Schellinx, “A new deconstruc- tive logic: Linear logic,” J. Symb. Log., vol. 62, no. 3, pp. 755–807, 1997.