• Aucun résultat trouvé

A program for the full axiom of choice

N/A
N/A
Protected

Academic year: 2022

Partager "A program for the full axiom of choice"

Copied!
23
0
0

Texte intégral

(1)

A program for the full axiom of choice

Jean-Louis Krivine

Université de Paris - C.N.R.S.

krivine@irif.fr December 15, 2021

Introduction

The Curry-Howard correspondence enables to associate a program with each proof in classi- cal natural deduction. But mathematical proofs not only use the rules of natural deduction, but alsoaxioms, essentially those of Zermelo-Frænkel set theory with axiom of choice. In or- der to transform these proofs into programs, you must therefore associate with each of these axioms suitable instructions, which is far from obvious.

The theory ofclassical realizability (c.r.) solves this problem for all axioms of ZF, by means of a very rudimentary programming language : theλc-calculus, that is to say theλ-calculus with a control instruction [6].

The programs obtained in this way can therefore be written in practically any programming language. They are said torealizethe axioms of ZF.

But, almost all the applications of mathematics in physics, probability, statistics, etc. use Analysis, that is to saythe axiom of dependent choice. The first program for this axiom, known since 1998 [1], is a pureλ-term calledbar recursion[20, 2, 21]. In fact, c.r. shows that it pro- vides also a program for thecontinuum hypothesis[17].

Nevertheless, this method requires the programming language to be limited toλc-calculus, prohibiting any other instruction, which is a severe restriction.

Classical realizability provides other programs for the axiom of dependent choice [13, 14, 15], which require an additional instruction (clock, signature, . . . , or as in this paper, introduc- tion offresh variables). On the other hand, this kind of solution is very flexible regarding the programming language.

There remained, however, the problem of thefullaxiom of choice. It is solved here, by means of new instructions which allowbranching or parallelism1. Admittedly, the program ob- tained in this way is rather complicated and we may hope for new solutions as simple as bar recursion or clock. However, it already shows that we can turn all proofs of ZFC into pro- grams (more precisions about this in the remark at the end of the paper).

Note thatthe proof is constructivei.e. it gives explicitly a program for AC : you ”merely” need to formalize this proof itself in natural deduction and apply the proof-program correspon- dence to the result.

1A detailed study of the effect of such instructions on the characteristic Boolean algebraג2 is made in [5].

(2)

Outline of the paper.

The framework of this article is thetheory of classical realizability, which is explained in de- tail in [13, 14, 15]. For the sake of simplicity, we will often refer to these papers for definitions and standard notations.

Section 1. Axioms and properties ofrealizability algebras (r.a.) which are first order struc- tures, a generalization of both combinatory algebras and ordered sets of forcing conditions.

Intuitively, their elements are programs.

Section 2. Each r.a. is associated with arealizability model (r.m.), a generalization of the generic model in the theory of forcing. This model satisfies a set theory ZFεwhich is a con- servative extension of ZF, with a strong non extensional membership relationε.

Section 3.Generic extensions of realizability algebras and models. We explain how a generic extension of a given r.m. can be handled by means of a suitable extension of the correspond- ing r.a. A little technical because it mixes forcing and realizability.

Section 4. Definition and properties of a particular realizability algebra A0. It contains an instruction of parallelismand the programming language allows the introduction of many other instructions. This is therefore, in fact, a class of algebras.

Using this r.a., we obtain a program for theaxiom of well ordered choice (WOC) : the product of a family of non-empty sets whose index set is well ordered is non-empty.

We get, at the same time, a new proof that WOC is weaker than AC (joint work with Laura Fontanella ; cf. [8] for the usual proof ).

Section 5. Construction of a generic extensionA1of the algebraA0which allows torealize the axiom of choice : every set can be well ordered.

Sections 3 and 4 are independent.

Thanks to Asaf Karagila for several fruitful discussions [10]. He observed that, by exactly the same method, we obtain a program, not only for AC, but for every axiom which can be shown compatible with ZFC by forcing (theorem 33) ; for instance CH, 20= ℵ2+ Martin’s axiom, etc.

(but this does not apply to V = L).

1 Realizability algebras (r.a.)

It is a first order structure, which is defined in [14]. We recall here briefly this definition and some essential properties :

Arealizability algebraA is made up of three sets :Λ(the set ofterms),Π(the set ofstacks), Λ ⋆ Π(the set ofprocesses) with the following operations :

(ξ,η)7→(ξ)ηfromΛ2intoΛ(application) ; (ξ,π)7→ξ

.

πfromΛ×ΠintoΠ(push) ;

(ξ,π)7→ξπfromΛ×ΠintoΛ ⋆ Π(process) ; π7→kπfromΠintoΛ(continuation).

There are, inΛ, distinguished elements B,C,I,K,W,cc, calledelementary combinatorsorin- structions.

Notation.The term (. . . (((ξ)η12) . . .)ηnwill be also written as (ξ)η1η2. . .ηnorξη1η2. . .ηn. For instance : ξηζ=(ξ)ηζ=(ξη)ζ=((ξ)η)ζ.

(3)

We define a preorder onΛ ⋆ Π, denoted by ≻, which is calledexecution; ξπξπis read as :the processξπreduces toξπ.

It is the smallest reflexive and transitive binary relation, such that, for any ξ,η,ζ∈Λand π,ϖ∈Π, we have :

(ξ)ηπξη

.

π(push).

I⋆ξ

.

πξπ(no operation).

K⋆ξ

.

η

.

πξπ(delete).

Wξ

.

η

.

πξη

.

η

.

π(copy).

C⋆ξ

.

η

.

ζ

.

πξζ

.

η

.

π(switch).

B⋆ξ

.

η

.

ζ

.

πξ⋆(η)ζ

.

π(apply).

ccξ

.

πξ⋆kπ

.

π(save the stack).

kπξ

.

ϖξπ(restore the stack).

We are also given a subset⊥⊥ofΛ ⋆ Π, called ”the pole”, such that : ξπξπ,ξπ∈ ⊥⊥ ⇒ ξπ∈ ⊥⊥.

Given two processesξπ,ξπ, the notationξπ≻≻ξπmeans : ξπ∉ ⊥⊥ ⇒ξπ∉ ⊥⊥.

Therefore, obviously, ξπξπξπ≻≻ξπ.

Finally, we choose a set of terms PLA ⊂Λ, containing the elementary combinators :

B,C,I,K,W,ccand closed by application. They are called theproof-like terms of the algebraA. We write alsoPLinstead ofPLA if there is no ambiguity aboutA.

The algebraA is calledcoherent if, for every proof-like term θ∈PLA, there exists a stackπ such thatθπ∉ ⊥⊥.

Remark. Aset of forcing conditionscan be considered as a degenerate case of realizability algebra, if we present it in the following way : an inf-semi-latticeP, with a greatest element1and an initial segment⊥⊥ofP (the set offalseconditions). Two conditionsp,qP are calledcompatibleif their g.l.b.pqis not in⊥⊥.

We get a realizability algebra if we setΛ=Π=Λ ⋆ Π=P ;B=C=I=K=W=cc=1andPL={1} ; (p)q =p

.

q =pq=pq andkp=p. The preorder pq is defined aspq, i.e. pq=p. The condition of coherence is1∉ ⊥⊥.

We can translateλ-terms into terms ofΛbuilt with the combinatorsB,C,I,K,Win such a way that weak head reduction is valid :

λx t[x]⋆u

.

πt[u/x]π whereλx t[x],uare terms andπis a stack.

This is done in [14]. Note that the usual (K,S)-translation does not work.

λ-calculus is much more intuitive than combinatory algebra in order to write programs, so that we use it extensively in the following. But combinatory algebra is better for theory, in particular because it is a first order structure.

2 Realizability models (r.m.)

The framework is very similar to that of forcing, which is anyway a particular case.

We use a first order language with three binary symbolsε/ ,∉,⊂(εis intended to be a strong,

(4)

non extensional membership relation ;∈and⊂have their usual meaning in ZF).

First order formulas are written with the only logical symbols→,∀,⊥,⊤.

The symbols¬,∧,∨,↔,∃are defined with them in the usual way.

Given a realizability algebra, we get arealizability model (r.m.)as follows :

We start with a modelM of ZFC (or even ZFL) called theground model. The axioms of ZFL are written with the sublanguage {∉,⊂}.

We build a modelN of a new set theory ZFε, in the language {ε/ ,∉,⊂}, the axioms of which are given in [13]. We recall them below, using the following rather standard abbreviations : F1→(F2→. . . (FnG) . . .) is writtenF1,F2, . . . ,FnGor even⃗FG.

We use the notation∃x{F1,F2, . . . ,Fn} for∀x(F1→(F2→ · · · →(Fn→ ⊥)· · ·))→ ⊥.

Of course,xεyandxyare the formulas/y→ ⊥andxy→ ⊥.

The notationx=yF meansxy,yxF. Thusx=y, which represents the usual (extensional) equality of sets, is the pair of formulas {x⊂y,yx}.

We use the notations (∀xεa)F(x) for∀x(¬F(x)→/a) and (∃xεa)F(x) for ¬∀x(⃗F(x)→/a).

For instance, (∃xεy)(t=u) is the formula¬∀x(tu,ut/y).

The axioms of ZFεare the following : 0. Extensionality axioms.

xy(xy↔(∃zεy)x=z) ;xy(xy↔(∀zεx)zy).

1. Foundation scheme.

∀⃗a(x((yεx)F(y,⃗a)F(x,⃗a))→ ∀x F(x,⃗a)) for every formulaF(x,a1, . . . ,an).

The intuitive meaning of these axioms is thatεis a well founded relation and the relation∈ is obtained by “ collapsing ”εinto an extensional relation.

The following axioms essentially express that the relation εsatisfies the axioms of Zermelo- Fraenkelexcept extensionality.

2. Comprehension scheme.

∀⃗a∀x∃y∀z(zεy↔(zεxF(z,⃗a)) for every formulaF(z,⃗a).

3. Pairing axiom.

abx{aεx,bεx}.

4. Union axiom.

ab(∀xεa)(yεx)yεb.

5. Power set axiom.

abx(yεb)z(zεy↔(zεazεx)).

6. Collection scheme.

∀⃗axy(uεx)(v F(u,v,a)→(∃vεy)F(u,v,a)) for every formulaF(u,v,a).

7. Infinity scheme.

∀⃗a∀x∃y{xεy, (∀uεy)(∃v F(u,v,a)→(∃vεy)F(u,v,a))} for every formulaF(u,v,a).

It is shown in [13] that ZFεis aconservative extension of ZF.

For each formula F(⃗a) of ZFε (i.e. written with ε/ ,∉,⊂) with parameters⃗a in the ground modelM we define, inM, a falsity value∥F(⃗a)∥which is a subset ofΠand a truth value

|F(⃗a)|which is a subset ofΛ. The notationt||−F(⃗a) (read “trealizesF(⃗a)” or “tforcesF(⃗a)”

in the particular case of forcing) meanst∈ |F(⃗a)|.

(5)

We set first|F(⃗a)| ={t ∈Λ; (∀π∈ ∥F(⃗a)∥)(t⋆π∈ ⊥⊥)} so that we only need to define∥F(⃗a)∥, which we do by induction onF :

1. Definition of∥aε/b∥:

/b∥ ={π∈Π; (a,π)b} ;∥⊥∥ =Π;∥⊤∥ = ;;

2. Definition of∥a⊂b∥and∥a∉b∥by induction on (rk(a)∪rk(b), rk(a)∩rk(b)) :

ab∥ =S

c{t

.

π∈Π;t||−cb, (c,π)∈a} ;

ab∥ =S

c{t

.

u

.

π∈Π;t||−ca,u||−ac, (c,π)∈b}.

3. Definition of∥F∥for a non atomic formulaF, by induction on the length :

∥F →F∥ ={t

.

π∈Π; t||−F,π∈ ∥F∥} ;

∥∀x F(x)∥ =S

aF(a)∥.

This notion of realizability has two essential properties given by theorems 1 and 2 below.

They are proved in [13].

Theorem 1(Adequacy lemma). ||−is compatible with classical natural deduction, i.e. : If t1, . . . ,tn,t areλc-terms such that t1:F1, . . . ,tn :Fnt :F in classical natural deduction, then t1||−F1, . . . ,tn||−Fnt||−F .

In particular, any valid formula is realized by a proof-like term.

Remark.The proof of theorem 1 uses only item 3 in the above definition of∥F∥. In other words, the values of∥F∥for atomic formulasFare arbitrary. This will be used in sections 3 and 5.

Theorem 2. The axioms of ZFεare realized by proof-like terms.

It follows thatevery closed formula which is consequence of ZFεand, in particular, every con- sequence of ZF, is realized by a proof-like term.

In the following, we shall simply say “the formula F is realized” instead of “realized by a proof-like term” and use the notation ||−F.

Theorem 2 is validfor every r.a.The aim of this paper is to realize the full axiom of choice AC in some particular r.a.suitable for programming.

Remark. Note that AC is realized for any r.a. associated with a set of forcing conditions (generic extension ofM). But in this case, there is only one proof-like term which is the greatest element1.

We define the strong (Leibnitz) equalitya =b by∀z(aε/z/z). It is trivially transitive and it is symmetric by comprehension. This equality satisfy the first order axioms of equal- ity∀xy(x= y →(F(x)→F(y))) (by comprehension scheme of ZFε) and is therefore the equality in the r.m.N .

Lemma 3.a=b∥ = ∥⊤ → ⊥∥ ={ξ

.

π;ξ∈Λ,π∈Π}if a̸=b ;

a=a∥ = ∥⊥ → ⊥∥ ={ξ

.

π;ξ∈Λ,π∈Π,ξπ∈ ⊥⊥}.

Letz={b}×Πso that∥bε/z∥ =Π. Ifa̸=bthen∥aε/z∥ = ;and therefore :

/z→/z∥ = ∥⊤ → ⊥∥.

Ifa=b, then/z∥ =Πand therefore∥a=b∥ = ∥⊥ → ⊥∥.

Q.E.D.

Finally, it is convenient to define first̸=by∥a̸=a∥ = ∥⊥∥ =Π;∥a̸=b∥ = ∥⊤∥ = ;ifa̸=b ; and to definea=basa̸=b→ ⊥.

(6)

We define a preorder≤on the setP(Π) of “falsity values” by setting :

XY ⇔ there exists a proof-like term θ||−XY. By theorem 1, we easily see [13] that (P(Π),≤) is a Boolean algebraBA if the r.a. A is coherent. Every formulaF(⃗a) of ZFε with parameters in the ground modelM has a value∥F(⃗a)∥in this Boolean algebra.

By means of any ultrafilter onBA, we thus obtain a complete consistent theory in the lan- guage {ε/ ,∉,⊂} with parameters inM. We take any modelN of this theory and call itthe realizability model (r.m.)of the realizability algebraA.

Therefore,N is a model of ZFε, and in particular, a model of ZF, that we will callN. ThusNis simply the modelN restricted to the language {∉,⊂}.

Remarks.

The ground modelM is contained inN since every element of it is a symbol of constant. ButM is not a submodel ofN for the common language {∉,⊂} ; and, except in the case of forcing, not every element ofN “has a name” inM.

WhenF is a closed formula of ZFε, the two assertionsN |=F and ||−F have essentially the same meaning, sinceN representsany r.m.for the given r.a. But the second formulation requires a formal proof.

Functionals

Afunctional relation defined inN is given by a formulaF(x,y) of ZFεsuch that : N |= ∀xyy(F(x,y),F(x,y)→y=y).

Afunctionis a functional relation which is a set.

We define now some special functional relations on N which we callfunctionals defined inM orfunctional symbols:

For each functional relationf :Mk→M defined in the ground modelM, we add the func- tional symbolf to the language of ZFε. The application off to an argumentawill be denoted

f[a]. Thereforef is also defined inN .

We call this (trivial) operationthe extension toN of the functional f defined in the ground model. It is a fundamental tool in all that follows.

Note the use of brackets for f[a] in this case.

Theorem 4. Let t1,u1. . . ,tn,un,t,u be k-ary terms built with functional symbols, such that : M |= ∀⃗x(t1[⃗x]=u1[⃗x], . . . ,tn[⃗x]=un[⃗x]t[x]=u[x]).

Thenλx1. . .λxnλx(x1) . . . (xn)x||− ∀⃗x(t1[⃗x]=u1[⃗x], . . . ,tn[⃗x]=un[⃗x]t[⃗x]=u[⃗x]).

This easily follows from the definition above of∥a=b∥.

Q.E.D.

As a first example let the unary functionalΦF[X] be defined inM by : ΦF[X]={(x,ξ

.

π) ;ξ||−F(x), (x,π)∈X}.

We shall denote it by {xεX ; F(x)} (in this notation,xis a bound variable) because it corre- sponds to the comprehension scheme in the modelN . Note that the use ofεreminds that this expression must be interpreted inN .

We define now inM the unary functionalגX =X×Π, so that we have :

∥xεX∥ =ΠifxX and∥xε/גX∥ = ;ifxX.

For anyX inM, we define the quantifier∀xגX by setting∥∀xגXF(x)∥ =S

xXF(x)∥.

(7)

Lemma 5. ||− ∀xגXF(x)↔ ∀x(xεגXF(x)).

In fact, we have∥∀x(¬F(x)→X)∥ = ∥∀xגX¬¬F(x)∥. Now we have triviallyλx(x)I||− ∀xגXF(x)→ ∀xגX¬¬F(x) andcc||− ∀xגX¬¬F(x)→ ∀xגXF(x).

Q.E.D.

Lemma 6. Let f be a functional k-ary symbol defined inM such that f :X1×· · ·×XkX . Then its extension toN is such that fX1×· · ·×גXk→גX .

Trivial.

Q.E.D.

By theorem 4 and lemma 6, the algebra operations on the Boolean algebra 2={0, 1} extended toג2, turn it into a Boolean algebra which we callthe characteristic Boolean algebraof the r.m.N. In the ground modelM, we define the functional (a,x)7→axfrom 2×M intoM by 0x= ;and 1x=x. It extends toN into a functionalג2×N N such that (ab)x=a(bx) for a,bεג2 and everyxinN .

Lemma 7.

I||− ∀⃗x∀⃗yaג2(a f[⃗x,⃗y]=a f[a⃗x,⃗y])for every functional symbol f defined inM. Immediate sinceξ||− ∀aג2F(a) means (ξ||−F(0))∧(ξ||−F(1)).

Q.E.D.

For any formulaF(⃗x) of ZF we define, inM, a functional〈F(⃗x)〉with value in {0, 1} which is the truth value of this formula inM.2The extension of this functional to the modelN takes its values in the Boolean algebraג2 (cf. [16]).

The binary functionals 〈x ∉ y〉 and〈x ⊂ y〉 define on the r.m. N a structure of Boolean modelon the Boolean algebraג2, that we denote byMג2. It is an elementary extension ofM since the truth value of every closed formula of ZF with parameters inM is the same inM andMג2.

Any ultrafilterU onג2 would therefore give a (two-valued) modelMג2/U which is an ele- mentary extension ofM. In [16], it is shown that there exists one and only one ultrafilterD onג2 such that the modelMג2/D, which we shall denote asMD, is well founded (inN ). The binary relations∉,⊂ofMDare thus defined by〈xyεDand〈xyεD.

Moreover,MD is isomorphic with a transitive submodel ofN with the same ordinals. In fact, if we start with a ground modelM which satisfies V = L, thenMD is isomorphic with the constructible class ofN.

Remark.We have definedfour first order structureson the modelN :

•The realizability modelN itself uses the language {ε/ ,∉,⊂} of ZFε; the equality onN is the Leibnitz equality =, which is the strongest possible.

•The modelNof ZF is restricted to the language {∉,⊂} ; the equality onNis the extensional equal- ity =.

• The Boolean modelMג2with the language {∉,⊂} of ZF and with truth values inג2 ; it is an ele- mentary extension of the ground modelM. The equality onMג2is〈x=y〉 =1 which is the same as Leibnitz equality.

2The formal definition in ZF is :∀⃗x((F(x)→ 〈F(x)〉 =1)(¬F(x)→ 〈F(x)〉 =0)).

(8)

• The modelMD with the same language, also an elementary extension ofM ; ifF(a) is a closed formula of ZF with parameters (inN), thenMD|=F(⃗a) iffN |= 〈F(⃗a)〉εD.

The equality onMDis given by〈x=y〉εD.

The proof of existence of the ultrafilter D in [16] is not so simple. But it is useless in the present paper, becauseג2 will be the four elements algebra, with two atomsa0,a1which give the two trivial ultrafilters onג2. It is easily seen (lemma 23) that one of them, saya0gives a well founded model denoted byMa0 which is the classa0N =MD. The classMa1=a1N is also an elementary extension ofM (but not well founded, cf. the remark after lemma 29).

Finally we haveMג2=N =a0a1N since the Boolean modelMג2is simply a product in this case, and equality is the same onMג2andN , as remarked previously.

Integers

We define, in the ground modelM, the functional x 7→ x+ =x∪{x} and extend it to the r.m.N . It is injective inM and therefore also inN .

For eachn∈N, we definen∈PLby induction : 0=λxλy y=KI;n+=sn wheres=λnλfλx(n f)(f)x=(BW)(B)B(andn+isn+1).

We defineNe={(n,n

.

π) ; n∈N,π∈Π}. We can use it as the set of integers of the modelN as shown by the following theorem 8.

We define the quantifier∀nintby setting∥∀nintF(n)∥ ={n

.

π;n∈N,π∈ ∥F(n)∥}.

Theorem 8. For every formula F(x)of ZFε, the following formulas are realized : i) 0εNe;∀n(nεNe→n+εN)e ;

ii) F(0),∀n(F(n)→F(n+))→(∀N)Fe (n); iii)nintF(n)↔(∀Ne)F(n).

i) Letξ||−0ε/Ne; thenξ⋆0

.

π∈ ⊥⊥for allπ∈Π. Thereforeλx(x)0||−0ε/Ne→ ⊥.

Letξ||−n+ε/Ne andn

.

π∈ ∥nε/N∥. We havee ξsn

.

π∈ ⊥⊥and thereforeθξ

.

n

.

π∈ ⊥⊥with θ=λxλn(x)(s)n. Thusθ||− ∀n(n+ε/Ne→/Ne).

ii) We showλxλyλzλn(x)(n)y z||− ¬F(0),∀n(F(n+)→F(n))→ ∀m(F(m)→/Ne).

Letξ||− ¬F(0),η||− ∀n(F(n+)→F(n)),ζ||−F(m) andm

.

π∈ ∥mε/N∥.e Let us show thatmηζ||−F(0) by induction onm:

This is clear ifm=0. Now,m+ηζ=smηζ≻(mη)(η)ζandηζ||−F(m) sinceζ||−F(m+).

Therefore (mη)(η)ζ||−F(0) by the induction hypothesis.

It follows that (ξ)(m)ηζ||− ⊥, hence the result.

iii) Let us use lemma 9. We have :

∥∀n(¬F(n)→/N)∥ =e {kπ

.

n

.

ϖ;n∈N,π∈ ∥F(n)∥,ϖ∈Π}

and by definition :∥∀nintF(n)∥ ={n

.

π;n∈N,π∈ ∥F(n)∥}.

It follows easily that :

λxλn(cc)λk(x)kn||− ∀n(¬F(n)→/N)e → ∀nintF(n) λxλkλn(k)(x)n||− ∀nintF(n)→ ∀n(¬F(n)→/Ne).

Q.E.D.

Some useful notations

For every set of termsX ⊂Λand every closed formulaF we can define an “extended formula”

XF by setting∥XF∥ ={ξ

.

π;ξX,π∈ ∥F∥}.

(9)

For instance, for every formulaF, we define¬F ={kπ; π∈ ∥F∥}. It is a useful equivalent of

¬F by the following : Lemma 9. ||−¬F ↔ ¬F .

Ifπ∈ ∥F∥, thenkπ||− ¬F and thereforeI ||−¬F→ ¬F.

Conversely, ifξ||−¬F → ⊥, thenξ⋆kπ

.

π∈ ⊥⊥for everyπ∈ ∥F∥, thereforecc||− ¬¬FF.

Q.E.D.

Ift,u are terms of the language of ZF, built with functionals inM, we define another “ex- tended formula”t=u,→F by setting :

∥t=u,→F∥ = ;ift̸=u;∥t=u,→F∥ = ∥F∥ift=u.

We write (t1=u1), . . . , (tn=un),→F for (t1=u1),→(· · ·,→((tn=un),→F)· · ·).

Lemma 10. ||−((t=u,→F)↔(t=uF)).

We have immediately I ||− ¬F, (t=u,→F)→t̸=u.

Converselyλx(x)I||−(t=uF)→(t=u,→F).

Q.E.D.

For instance, the conclusion of theorem 4 may be rewritten as : I ||− ∀⃗x(t1[⃗x]=u1[⃗x], . . . ,tn[⃗x]=un[⃗x],→t[⃗x]=u[⃗x]).

Lemmas 11, 12, 13 and theorem 15 below will be used in the following sections.

Lemma 11. ||− ∀x∀y(xεy→ 〈x∈Cl[y]〉 =1).

In the modelM, the unary functional symbol Cl denotes thetransitive closure.

We show I ||− ∀xy(x∈Cl[y]〉 ̸=1→/y) : letξ||− 〈x∈Cl[y]〉 ̸=1 andπ∈ ∥/y∥. Then (x,π)y, thereforex∈Cl(y). It follows thatξ||− ⊥.

Q.E.D.

Lemma 12. I ||− ∀Xaג2x(axa X〉 ≥ a ,→ axεΦ[X]) whereΦ is the functional symbol defined inM byΦ[X]=(X∪{0})×Πi.e.ג(X ∪{0}).

Remark.The notationbafora,bεג2 means, of course,ab=a. This amounts to show :

1. xX ⇒I||−/Φ[X]→ ⊥; 2. I||−0ε/Φ[X]→ ⊥.

Both are trivial.

Q.E.D.

Lemma 13. Let F(x,⃗y)be a formula in ZFε. Then :

I||− ∀⃗y(∀ϖגΠF(f[ϖ,⃗y],⃗y)→ ∀x F(x,⃗y)) for some functional symbol f defined inM.

Let⃗a=(a1, . . . ,ak) inM. Then, we have :

π∈ ∥∀x F(x,⃗a)∥ ⇔M |= ∃x(π∈ ∥F(x,⃗a)∥)⇔M |=π∈ ∥F(f[π,⃗a],⃗a)∥ wheref is a functional defined inM, (choice principle inM).

Thus we have∥∀x F(x,⃗a)∥ ⊂S

ϖ∈ΠF(f[ϖ,⃗a],a)∥hence the result.

Q.E.D.

(10)

Lemma 14.

If X∈M is transitive, thenגX is ε-transitive, i.e. I||− ∀xגX∀y(yεX yε/x).

LetxX,ϖ∈ ∥yε/x∥, i.e. (y,ϖ)xandξΛsuch thatξ||−X. SinceX is transitive, we haveyX and thereforeξ||− ⊥.

Q.E.D.

Theorem 15. LetLbe the language{ε,∈,⊂}of ZFε, with a symbol for each functional definable inM. Then, there exists anε-transitiveL-elementary substructureNfofN such that, for all a inNf, there is an ordinalαofM such thatNf|=גVα.

Nfis made up of the elementsaofN such thatגVαfor some ordinalαofM (note that it is not a class defined inN). By lemma 14, eachגVαisε-transitive and thereforeNfis also.

LetF(x,⃗y) be a formula ofLand⃗b=(b1, . . . ,bk) be inNf. We assumeNf|= ∀x F(x,⃗b) and we have to show thatN |= ∀x F(x,⃗b) which we do by induction onF. By lemma 13, it suffices to show thatN |= ∀ϖגΠF(f[ϖ,⃗b],b). Thus letπεגΠ; by definition ofNf, there existsX inM such that⃗גX. Now, there existsαinM such thatf :Π×XVαand therefore, inN, we have f :גΠ×גX →גVα. It follows thatf[π,⃗b]εגVαand therefore f[π,⃗b] is inNf.

ThusNf|=F(f[π,⃗b],b), henceN |=F(f[π,⃗b],b) by the induction hypothesis.

Q.E.D.

ReplacingN by this elementary substructure, we shall suppose from now on :

11

For all a inN , there is an ordinalαofM such thatN |=גVα.

Theorem 16 below is not really useful in the following, but it gives a welcome information on the (very complex) structure of the r.m.N.

Theorem 16. aN is a proper class for all aεג2,a̸=0; i.e. I||− ∀x(∀aג2(∀y(a yεx)a=0)).

Let us show, in fact, that I||− ∀x(∀aג2(axεxa=0)) ; in other words : I||−0εx→0=0 and I||−xεx→1=0.

We have (0εx)≡(0ε/x→ ⊥) and (0=0)≡(⊥ → ⊥) hence the first result.

Furthermore, we have∥xε/x∥ ={π∈Π; (x,π)x}= ;; thus∥xεx∥ = ∥xε/x→ ⊥∥ = ∥⊤ → ⊥∥

and∥1=0∥ = ∥⊤ → ⊥∥hence the second result.

Q.E.D.

3 Extensional generic extensions

In this section we build some tools in order to manage generic extensionsN[G] of theex- tensionalmodelN. We define a new r.a. and give, in this r.a., a new way to compute the truth value of ZF-formulas inN[G].

LetVbe fixed inM. We have∥V∥ = ∥〈xV〉 ̸=1∥and therefore : I||− ∀x(xεגV↔ 〈xV〉 =1).

In other words, a setxstrongly belongs toגViff the Boolean ZF-modelMג2says that “xbe- longs always toV”.

Remember also the important (and obvious) equivalence I||− ∀x∀y(x=y↔ 〈x=y〉 =1) which follows from ∥x ̸= y∥ = ∥〈x = y〉 ̸= 1∥ and identifies the r.m. N with the Boolean modelMג2.

(11)

Theorem 17. N |= ∀anint³

V)n→ ∃!bג(V

n)iint(i<na(i)=b[i])

´ .

In other words, each finite sequence of גVinN is represented by a unique finite sequence of V in the Boolean modelMג2.

Unicity. Note first that, since there is no extensionality inN , you may have two sequences a̸=aε(גV)nsuch thata(i)=a(i) fori<n. Now supposeb,bεג(Vn) be such thatb[i]=b[i] fori <n. Then, we haveb,bVn〉 =1 and〈(∀i <n)(b[i]=b[i])〉 =1. Since the Boolean modelMג2satisfies extensionality, we get〈b=b〉 =1 that isb=b.

Existence. Proof by induction onn. This is trivial ifn=0 : takeb= ;.

Now, letV)n+1andabe a restriction ofa ton. Letbεג(Vn) such that a(i)=b[i] for i<n(induction hypothesis).

In the ground modelM, we define the binary functional +as follows :

ifuis a finite sequence (u0, . . . ,un1), thenu+v is the sequence (u0, . . . ,un1,v).

We extend it toN and we setb=b+a(n) i.e.〈b=b+a(n)〉 =1. Therefore〈b[i]=b[i]〉 =1 fori<nand〈b[n]=a(n)〉 =1, i.e.a(i)=b[i] fori<nanda(n)=b[n].

Q.E.D.

Consider an arbitrary ordered set (C,≤) in the model N. By theorem 15 and property ⃝⃝11

(section 2), we may suppose that (C,≤)εגVwithV=VαinM.גVisε-transitive, by lemma 14.

As a set of forcing conditions,C is equivalent to the setCof finite subsetsX ofגVsuch that XC has a lower bound inC,Cbeing ordered by inclusion.

Thus, we can defineCby the following formula of ZFε:

C(u)≡uε(גV)(Im(u)∩C)has a lower bound in C where Im(u)⊆גVis the (finite) image of the finite sequenceu.

We haveN |=C(u)→V)and thereforeN |=C(u)→V) by theorem 17.

Moreover, we haveN |=C(;).

InM, the function (u,v)7→ uv, from (V)2 into V, which is the concatenation of se- quences, is associative with;as neutral element, also denoted by1(monoïd).

This function extends toN into an application ofג(V)2intoג(V) with the same proper- ties. Thus, we writeuv wforu(v w), (uv)w, etc.

The formulaC(uv) of ZFεmeans thatu,v are two compatible finite sequences of elements ofC, i.e. the union of their images has a lower bound inC. ThusCbecomes a set of forcing conditions equivalent toC by means of this compatibility relation.

This formula has the following properties :

||−C(puv q)→C(p vuq),||−C(puv q)→C(puq),||−C(puq)→C(puuq).

It will be convenient to have only one formula and to use simply the following consequence :

22

There exists a proof-like termcsuch that c||−C(pqr t uv w)→C(pt r uuv).

Consider now, in the ground modelM, a r.a.A0which gives the r.m.N . We suppose to have an operation (π,τ)7→πτfromΠ×ΛintoΠsuch that :

(ξ

.

π)τ=ξ

.

πτfor everyξ,τ∈Λandπ∈Π. and two new combinatorsχ(read)andχ(write)such that :

χξ

.

πτ≻≻ξτ

.

π(i.e.ξτ

.

π∈ ⊥⊥ ⇒χξ

.

πτ∈ ⊥⊥) χτ

.

ξ

.

π≻≻ξπτ(i.e. ξπτ∈ ⊥⊥ ⇒χτ

.

ξ

.

π∈ ⊥⊥).

Moreover, we suppose thatχ,χmay be used to form proof-like terms.

Intuitively,πτis obtained by putting the termτat the endof the stackπ, in the same way that τ

.

πis obtained by puttingτat the topofπ.

(12)

We define now, in the ground modelM, a new r.a. A1; its r.m. will be calledthe extension ofN by aC-generic (or a C -generic).

We define the termsB,C,I,K,W,ccandkπby the conditions :

33

B=B;I=I;

Cξ

.

η

.

ζ

.

πτξζ

.

η

.

πcτ; i.e.C=λxλyλz(χ)λt((χ)(c)t)xz y; Kξ

.

η

.

πτξπcτ; i.e.K=λxλy(χ)λt((χ)(c)t)x;

Wξ

.

η

.

πτξη

.

η

.

πcτ; i.e.W=λxλy(χ)λt((χ)(c)t)x y y; kπξ

.

ϖτξπcτ; i.e.kπ=λx(χ)λt(kπ)((χ)(c)t)x;

ccξ

.

πτξ⋆kπ

.

πcτ; i.e.

cc=λx(χ)λt(cc)λk(((χ)(c)t)x)λx(χ)λt(k)((χ)(c)t)x.

When checking below the axioms of r.a., the property needed for each combinator is :

44

forC:c||−C(pr t v)→C(p t r v) ; forK:c||−C(pqr)→C(pr) ; forW:c||−C(puv)→C(puuv) ; forcc:c||−C(pu)→C(puu) ; forkπ:c||−C(r t w)→C(t r).

We get them replacing by1some of the variablesp,q,r,t,u,v,win the definition⃝⃝22 ofc. We define the r.a.A1by settingΛ=Λ×V;Π=Π×V;ΛΠ=(Λ ⋆ Π)×V.

(ξ,u)

.

(π,v)=(ξ

.

π,uv) ; (ξ,u)⋆(π,v)=(ξπ,uv) ; (ξ,u)(η,v)=(ξη,uv).

The pole⊥⊥⊥ofA1is defined by :

(ξ⋆π,u)∈ ⊥⊥⊥ ⇔(∀τ∈Λ)(τ||−C(u)⇒ξπτ∈ ⊥⊥).

The combinators are :

B=(B,1),C=(C,1),I=(I,1),K=(K,1),W=(W,1),cc=(cc,1) andk(π,u)=(kπ,u).

PLA1 is {(θ,1) ;θ∈PLA0}.

Theorem 18. A1is a coherent r.a.

Let us prove first thatA1is coherent : letθ∈PLA0; by definition of the formulaC, there exists τ0∈PLA0such thatτ0||−C(1).

Sinceχθτ0∈PLA0, there exists π∈Πsuch thatχθτ0π∉ ⊥⊥, thusθπτ0 ∉ ⊥⊥, therefore (θ,1)⋆(π,1)∉ ⊥⊥⊥.

We check now thatA1is a r.a. :

• (ξ,u)⋆(η,v)

.

(π,w)∈ ⊥⊥⊥ ⇒(ξ,u)(η,v)⋆(π,w)∈ ⊥⊥⊥.

By hypothesis, we have (ξ⋆η

.

π,uv w)∈ ⊥⊥⊥, therefore (∀τ∈Λ)(τ||−C(uv w)→ξη

.

πτ∈ ⊥⊥) and therefore (∀τ∈Λ)(τ||−C(uv w)→ξηπτ∈ ⊥⊥) hence the result.

• (ξ,u)⋆(η,v)(ζ,w)

.

(π,z)∈ ⊥⊥⊥ ⇒(B,1)⋆(ξ,u)

.

(η,v)

.

(ζ,w)

.

(π,z)∈ ⊥⊥⊥.

By hypothesis, we have (ξηζ

.

π,uv w z)∈ ⊥⊥⊥therefore :

(∀τ∈Λ)(τ||−C(uv w z)→ξηζ

.

πτ∈ ⊥⊥) thus (∀τ∈Λ)(τ||−C(uv w z)→Bξ

.

η

.

ζ

.

πτ∈ ⊥⊥).

• (ξ,u)⋆(ζ,w)

.

(η,v)

.

(π,z)∈ ⊥⊥⊥ ⇒(C,1)⋆(ξ,u)

.

(η,v)

.

(ζ,w)

.

(π,z)∈ ⊥⊥⊥. By hypothesis, we have (ξ⋆ζ

.

η

.

π,uw v z)∈ ⊥⊥⊥, therefore :

(∀τ∈Λ)(cτ||−C(uw v z)→ξζ

.

η

.

πcτ∈ ⊥⊥) therefore, by definition⃝⃝33 ofC:

(∀τ∈Λ)(cτ||−C(uw v z)→Cξ

.

η

.

ζ

.

πτ∈ ⊥⊥). But, by the property⃝⃝44 ofc, we have : τ||−C(uv w z)→cτ||−C(uw v z), hence the result.

• (ξ,u)⋆(π,v)∈ ⊥⊥⊥ ⇒(I,1)⋆(ξ,u)

.

(π,v)∈ ⊥⊥⊥.

(13)

Immediate.

• (ξ,u)⋆(π,w)∈ ⊥⊥⊥ ⇒(K,1)⋆(ξ,u)

.

(η,v)

.

(π,w)∈ ⊥⊥⊥. By hypothesis, we have (ξ⋆π,uw)∈ ⊥⊥⊥, therefore :

(∀τ∈Λ)(cτ||−C(uw)→ξπcτ∈ ⊥⊥), therefore by the definition⃝⃝33 ofK: (∀τ∈Λ)(cτ||−C(uw)→Kξ

.

η

.

πτ∈ ⊥⊥). But, by⃝⃝44 , we have :

τ||−C(uv w)→cτ||−C(uw), hence the result.

• (ξ,u)⋆(η,v)

.

(η,v)

.

(π,w)∈ ⊥⊥⊥ ⇒(W,1)⋆(ξ,u)

.

(η,v)

.

(π,w)∈ ⊥⊥⊥. By hypothesis, we have (ξ⋆η

.

η

.

π,uv v w)∈ ⊥⊥⊥, therefore :

(∀τ∈Λ)(cτ||−C(uv v w)→ξη

.

η

.

πcτ∈ ⊥⊥), thus, by the definition⃝⃝33 ofW: (∀τ∈Λ)(cτ||−C(uv v w)→Wξ

.

η

.

πτ∈ ⊥⊥). Now, by⃝⃝44, we have :

τ||−C(uv w)→cτ||−C(uv v w), hence the result :τ||−C(uv w)→Wξ

.

η

.

πτ∈ ⊥⊥.

• (ξ,v)⋆(π,u)∈ ⊥⊥⊥ ⇒(kπ,u)⋆(ξ,v)

.

(ϖ,w)∈ ⊥⊥⊥.

By hypothesis, we have (ξ⋆π,vu)∈ ⊥⊥⊥, therefore :

(∀τ∈Λ)(cτ||−C(vu)→ξπcτ∈ ⊥⊥), thus, by the definition⃝⃝33 ofkπ: (∀τ∈Λ)(cτ||−C(vu)→kπξ

.

ϖτ∈ ⊥⊥). Now, by⃝⃝44 , we have :

τ||−C(uv w)→cτ||−C(vu), hence the result :τ||−C(uv w)→kπξ

.

ϖτ∈ ⊥⊥.

• (ξ,u)⋆(kπ,v)

.

(π,v)∈ ⊥⊥⊥ ⇒(cc,1)⋆(ξ,u)

.

(π,v)∈ ⊥⊥⊥.

By hypothesis, we have (ξ⋆kπ

.

π,uv v)∈ ⊥⊥⊥, therefore :

(∀τ∈Λ)(cτ||−C(uv v)→ξ⋆kπ

.

πcτ∈ ⊥⊥), thus, by the definition⃝⃝33 ofcc: (∀τ∈Λ)(cτ||−C(uv v)→ccξ

.

πτ∈ ⊥⊥). But, by⃝⃝44 , we have :

τ||−C(uv)→cτ||−C(uv v), hence the result :τ||−C(uv)→ccξ

.

πτ∈ ⊥⊥.

Q.E.D.

The C-forcing defined in N N

LetF(⃗a) be a closed formula of the language of ZF with parameters inM.

We define the formulapF(⃗a) (read “p forces F(⃗a)”)as the formula of ZFεwhich expresses the C -forcing onN. In this formula, the variablepis restricted toV.

We define a subset (((F(⃗a)))) ofΠby (((F(⃗a))))={(π,p) ;π∈ ∥p⊮¬F(⃗a)∥}.

Lemma 19.

For every formula F(⃗x)of ZF, there exist two proof-like termspF,pF ofA such that : i) ξ||−(p⊩F(⃗a))⇒(pFξ,p)∥|−(((F(⃗a))))

ii) (ξ,p)∥|−(((F(⃗a))))⇒pFξ||−(p⊩F(⃗a)) for everyξ∈Λand pV.

By a well known property of forcing [9, 18], the formulapF(⃗x) is equivalent, in ZFε, to the formula∀q(C(p q)→q⊮¬F(⃗x)).

It follows that there are two proof-like termsqF,qF such that : (1) qF ||−(p⊩F(⃗a))→ ∀q(C(pq)→q⊮¬F(⃗a)) ;

(2) qF ||− ∀q(C(p q)→q⊮¬F(⃗a))→(p⊩F(⃗a)).

i) By applying (1) to the hypothesis, we obtainqFξ||− ∀q(C(pq)→q⊮¬F(⃗a)) that is :

∀π∀q∀τ(τ||−C(p q),π∈ ∥q⊮¬F(⃗a)∥ →qFξτ

.

π∈ ⊥⊥). It follows that :

∀π∀q(π∈ ∥q⊮¬F(⃗a)∥ → ∀τ(τ||−C(pq)→(χ)(qF)ξπτ∈ ⊥⊥)).

Nowπ∈ ∥q⊮¬F(⃗a)∥is the same as (π,q)∈(((F(⃗a)))) and

(14)

∀τ(τ||−C(p q)→(χ)(qF)ξπτ∈ ⊥⊥) is the same as (pFξ,p)⋆(π,q)∈ ⊥⊥⊥withpF =λx(χ)(qF)x.

Thus we obtain (pFξ,p)∥|−(((F(⃗a)))).

ii) The hypothesis gives∀π∀q((π,q)∈(((F(⃗a))))→(ξ⋆π,pq)∈ ⊥⊥⊥) that is :

∀π∀q∀τ(π∈ ∥q⊮¬F(⃗a)∥,τ||−C(p q)→ξπτ∈ ⊥⊥) and therefore :

∀π∀q∀τ(τ||−C(p q),π∈ ∥q⊮¬F(⃗a)∥ →χξτ

.

π∈ ⊥⊥).

It follows thatχξ||− ∀q(C(pq)→q⊮¬F(⃗a)) and therefore, by (2) : (qF)(χ||−(p⊩F(⃗a)). We setpF =λx(qF)(χ)x.

Q.E.D.

In the general theory of classical realizability, we define a truth value for the formulas of ZFε and therefore, in particular, for the formulas of ZF. We will define here directly a new truth value∥|F(a1, . . . ,an)∥|for a formula of ZFwith parameters inM for the r.a.A1.

To this aim, we first define the truth values∥|ab∥|,∥|ab∥|of the atomic formulas of ZF ; then that ofF(a1, . . . ,an), by induction on the length of the formula.

Theorem 1 (adequacy theorem) remains valid (cf. the remark after theorem 1) :

∥|a∉b∥| = (((a∉b))); ∥|a⊂b∥| =(((a⊂b))) ;

∥|FF∥| ={(ξ

.

π,p q) ; (ξ,p)∥|−F, (π,q)∈ ∥|F∥|} ;

∥|∀x F(x,a1, . . . ,an)∥| =S

a∥|F(a,a1, . . . ,an)∥|.

Of course ∥|−is defined by : (ξ,p)∥|−F ⇔(∀(π,q)∈ ∥|F∥|)((ξ,p)⋆(π,q)∈ ⊥⊥⊥).

Remark.Be careful, as we said before, these are not the truth values, in the r.a.A1, ofabandab considered as formulas of ZFε. We seek here to define directly theC-generic model onNwithout going through a model of ZFε.

Theorem 20 below may be considered as a generalization of the well known result about iteration of forcing : the r.a.A1, which is a kind of product ofA0byC, gives the same r.m. as theC-generic extension ofN.

Theorem 20.

For each closed formula F of ZF with parameters in the modelN , there exist two proof-like termsχF,χF, which only depend on the propositional structure of F , such that we have, for every ξ∈Λand pV:

ξ||−(p⊩F) ⇒ (χFξ,p)∥|−F and (ξ,p)∥|−FχFξ||−(p⊩F).

Thepropositional structureofF is the propositional formula built with the connective→and only two atomsO,O, which is obtained fromF by deleting all quantifiers and by identify- ing all atomic formulastu,turespectively withO,O.

For instance, the propositional structure of the formula :

X(∀x(y((x,y)∉YyX)→xX)→ ∀x(xX)) is ((OO)→O)→O. Proof by recurrence on the length ofF.

• IfF is atomic, we haveFaborab. Apply lemma 19 withF(⃗a)aborF(⃗a)ab.

• IfF ≡ ∀x F, thenpF≡ ∀x(p⊩F). Thereforeξ||−pF ≡ ∀x(ξ||−(p⊩F)).

Moreover, (ξ,p)∥|−F ≡ ∀x((ξ,p)∥|−F).

The result is immediate, from the recurrence hypothesis.

• IfFFF′′, we have pF ≡ ∀q(qFpqF′′) and therefore :

55

ξ||−(p⊩F)⇒ ∀η∀q(η||−(q⊩F)→ξη||−(pq⊩F′′)).

Suppose that ξ||−(p⊩F) and put χF=λxλyF′′)(x)(χF)y.

Références

Documents relatifs

Russell used acquaintance to that effect: ÒI say that I am acquainted with an object when I have a direct cognitive relation to that object, that is when I am directly aware of

Nevertheless, we have shown that it is sufficient for the modeling of the biological process at various levels in systems biology : our paradigmatic example of the switch of the λ

Now, there are essentially two possible axiom systems for mathematics : 1. second order classical logic with dependent choice... 2. Zermelo-Fraenkel set theory with the full axiom

QED Now, if we prove a formula © using the recurrence axiom, we know that the restricted formula © Int is provable without it, using formulas (*).. Therefore : If © is provable

It follows that the constructible universe of a realizability model is an elementary extension of the constructible universe of the ground model (a trivial fact in the particular

Hence since the birth of functional programming there have been several theoretical studies on extensions of the λ-calculus in order to account for basic algebra (see for

Nevertheless, it implies the dependant axiom of choice AC D (and in fact the non extensional axiom of choice AC N E ) and we will give and study the behaviour of the program

The crucial question to be addressed concerns the interpretation of the dS/AdS UIR’s (or quantum AdS and dS elementary systems) from a (asymptotically) minkowskian point of view..