HomeHome New Foundations Explorer
Theorem List (p. 14 of 64)
< Previous  Next >
Bad symbols? Try the
GIF version.

Mirrors  >  Metamath Home Page  >  NFE Home Page  >  Theorem List Contents       This page: Page List

Theorem List for New Foundations Explorer - 1301-1400   *Has distinct variable group(s)
TypeLabelDescription
Statement
 
Theoremnanbi1d 1301 Introduce a right anti-conjunct to both sides of a logical equivalence. (Contributed by SF, 2-Jan-2018.)
(φ → (ψχ))       (φ → ((ψ θ) ↔ (χ θ)))
 
Theoremnanbi2d 1302 Introduce a left anti-conjunct to both sides of a logical equivalence. (Contributed by SF, 2-Jan-2018.)
(φ → (ψχ))       (φ → ((θ ψ) ↔ (θ χ)))
 
Theoremnanbi12d 1303 Join two logical equivalences with anti-conjunction. (Contributed by Scott Fenton, 2-Jan-2018.)
(φ → (ψχ))    &   (φ → (θτ))       (φ → ((ψ θ) ↔ (χ τ)))
 
1.2.10  Logical 'xor'
 
Syntaxwxo 1304 Extend wff definition to include exclusive disjunction ('xor').
wff (φψ)
 
Definitiondf-xor 1305 Define exclusive disjunction (logical 'xor'). Return true if either the left or right, but not both, are true. After we define the constant true (df-tru 1319) and the constant false (df-fal 1320), we will be able to prove these truth table values: (( ⊤ ⊻ ⊤ ) ↔ ⊥ ) (truxortru 1358), (( ⊤ ⊻ ⊥ ) ↔ ⊤ ) (truxorfal 1359), (( ⊥ ⊻ ⊤ ) ↔ ⊤ ) (falxortru 1360), and (( ⊥ ⊻ ⊥ ) ↔ ⊥ ) (falxorfal 1361). Contrast with (df-an 360), (df-or 359), (wi 4), and (df-nan 1288) . (Contributed by FL, 22-Nov-2010.)
((φψ) ↔ ¬ (φψ))
 
Theoremxnor 1306 Two ways to write XNOR. (Contributed by Mario Carneiro, 4-Sep-2016.)
((φψ) ↔ ¬ (φψ))
 
Theoremxorcom 1307 is commutative. (Contributed by Mario Carneiro, 4-Sep-2016.)
((φψ) ↔ (ψφ))
 
Theoremxorass 1308 is associative. (Contributed by FL, 22-Nov-2010.) (Proof shortened by Andrew Salmon, 8-Jun-2011.)
(((φψ) ⊻ χ) ↔ (φ ⊻ (ψχ)))
 
Theoremexcxor 1309 This tautology shows that xor is really exclusive. (Contributed by FL, 22-Nov-2010.)
((φψ) ↔ ((φ ¬ ψ) φ ψ)))
 
Theoremxor2 1310 Two ways to express "exclusive or." (Contributed by Mario Carneiro, 4-Sep-2016.)
((φψ) ↔ ((φ ψ) ¬ (φ ψ)))
 
Theoremxorneg1 1311 is negated under negation of one argument. (Contributed by Mario Carneiro, 4-Sep-2016.)
((¬ φψ) ↔ ¬ (φψ))
 
Theoremxorneg2 1312 is negated under negation of one argument. (Contributed by Mario Carneiro, 4-Sep-2016.)
((φ ⊻ ¬ ψ) ↔ ¬ (φψ))
 
Theoremxorneg 1313 is unchanged under negation of both arguments. (Contributed by Mario Carneiro, 4-Sep-2016.)
((¬ φ ⊻ ¬ ψ) ↔ (φψ))
 
Theoremxorbi12i 1314 Equality property for XOR. (Contributed by Mario Carneiro, 4-Sep-2016.)
(φψ)    &   (χθ)       ((φχ) ↔ (ψθ))
 
Theoremxorbi12d 1315 Equality property for XOR. (Contributed by Mario Carneiro, 4-Sep-2016.)
(φ → (ψχ))    &   (φ → (θτ))       (φ → ((ψθ) ↔ (χτ)))
 
1.2.11  True and false constants
 
Syntaxwtru 1316 is a wff.
wff
 
Syntaxwfal 1317 is a wff.
wff
 
Theoremtrujust 1318 Soundness justification theorem for df-tru 1319. (Contributed by Mario Carneiro, 17-Nov-2013.)
((φφ) ↔ (ψψ))
 
Definitiondf-tru 1319 Definition of , a tautology. is a constant true. In this definition biid 227 is used as an antecedent, however, any true wff, such as an axiom, can be used in its place. (Contributed by Anthony Hart, 13-Oct-2010.)
( ⊤ ↔ (φφ))
 
Definitiondf-fal 1320 Definition of , a contradiction. is a constant false. (Contributed by Anthony Hart, 22-Oct-2010.)
( ⊥ ↔ ¬ ⊤ )
 
Theoremtru 1321 is provable. (Contributed by Anthony Hart, 13-Oct-2010.)
 
Theoremfal 1322 is refutable. (Contributed by Anthony Hart, 22-Oct-2010.) (Proof shortened by Mel L. O'Cat, 11-Mar-2012.)
¬ ⊥
 
Theoremtrud 1323 Eliminate as an antecedent. (Contributed by Mario Carneiro, 13-Mar-2014.)
( ⊤ → φ)       φ
 
Theoremtbtru 1324 If something is true, it outputs . (Contributed by Anthony Hart, 14-Aug-2011.)
(φ ↔ (φ ↔ ⊤ ))
 
Theoremnbfal 1325 If something is not true, it outputs . (Contributed by Anthony Hart, 14-Aug-2011.)
φ ↔ (φ ↔ ⊥ ))
 
Theorembitru 1326 A theorem is equivalent to truth. (Contributed by Mario Carneiro, 9-May-2015.)
φ       (φ ↔ ⊤ )
 
Theorembifal 1327 A contradiction is equivalent to falsehood. (Contributed by Mario Carneiro, 9-May-2015.)
¬ φ       (φ ↔ ⊥ )
 
Theoremfalim 1328 implies anything. (Contributed by FL, 20-Mar-2011.) (Proof shortened by Anthony Hart, 1-Aug-2011.)
( ⊥ → φ)
 
Theoremfalimd 1329 implies anything. (Contributed by Mario Carneiro, 9-Feb-2017.)
((φ ⊥ ) → ψ)
 
Theorema1tru 1330 Anything implies . (Contributed by FL, 20-Mar-2011.) (Proof shortened by Anthony Hart, 1-Aug-2011.)
(φ → ⊤ )
 
Theoremtruan 1331 True can be removed from a conjunction. (Contributed by FL, 20-Mar-2011.)
(( ⊤ φ) ↔ φ)
 
Theoremdfnot 1332 Given falsum, we can define the negation of a wff φ as the statement that a contradiction follows from assuming φ. (Contributed by Mario Carneiro, 9-Feb-2017.)
φ ↔ (φ → ⊥ ))
 
Theoreminegd 1333 Negation introduction rule from natural deduction. (Contributed by Mario Carneiro, 9-Feb-2017.)
((φ ψ) → ⊥ )       (φ → ¬ ψ)
 
Theoremefald 1334 Deduction based on reductio ad absurdum. (Contributed by Mario Carneiro, 9-Feb-2017.)
((φ ¬ ψ) → ⊥ )       (φψ)
 
Theorempm2.21fal 1335 If a wff and its negation are provable, then falsum is provable. (Contributed by Mario Carneiro, 9-Feb-2017.)
(φψ)    &   (φ → ¬ ψ)       (φ → ⊥ )
 
1.2.12  Truth tables

Some sources define operations on true/false values using truth tables. These tables show the results of their operations for all possible combinations of true () and false (). Here we show that our definitions and axioms produce equivalent results for (conjunction aka logical 'and') df-an 360, (disjunction aka logical inclusive 'or') df-or 359, (implies) wi 4, ¬ (not) wn 3, (logical equivalence) df-bi 177, (nand aka Sheffer stroke) df-nan 1288, and (exclusive or) df-xor 1305.

 
Theoremtruantru 1336 A identity. (Contributed by Anthony Hart, 22-Oct-2010.)
(( ⊤ ⊤ ) ↔ ⊤ )
 
Theoremtruanfal 1337 A identity. (Contributed by Anthony Hart, 22-Oct-2010.)
(( ⊤ ⊥ ) ↔ ⊥ )
 
Theoremfalantru 1338 A identity. (Contributed by Anthony Hart, 22-Oct-2010.)
(( ⊥ ⊤ ) ↔ ⊥ )
 
Theoremfalanfal 1339 A identity. (Contributed by Anthony Hart, 22-Oct-2010.)
(( ⊥ ⊥ ) ↔ ⊥ )
 
Theoremtruortru 1340 A identity. (Contributed by Anthony Hart, 22-Oct-2010.) (Proof shortened by Andrew Salmon, 13-May-2011.)
(( ⊤ ⊤ ) ↔ ⊤ )
 
Theoremtruorfal 1341 A identity. (Contributed by Anthony Hart, 22-Oct-2010.)
(( ⊤ ⊥ ) ↔ ⊤ )
 
Theoremfalortru 1342 A identity. (Contributed by Anthony Hart, 22-Oct-2010.)
(( ⊥ ⊤ ) ↔ ⊤ )
 
Theoremfalorfal 1343 A identity. (Contributed by Anthony Hart, 22-Oct-2010.) (Proof shortened by Andrew Salmon, 13-May-2011.)
(( ⊥ ⊥ ) ↔ ⊥ )
 
Theoremtruimtru 1344 A identity. (Contributed by Anthony Hart, 22-Oct-2010.)
(( ⊤ → ⊤ ) ↔ ⊤ )
 
Theoremtruimfal 1345 A identity. (Contributed by Anthony Hart, 22-Oct-2010.) (Proof shortened by Andrew Salmon, 13-May-2011.)
(( ⊤ → ⊥ ) ↔ ⊥ )
 
Theoremfalimtru 1346 A identity. (Contributed by Anthony Hart, 22-Oct-2010.)
(( ⊥ → ⊤ ) ↔ ⊤ )
 
Theoremfalimfal 1347 A identity. (Contributed by Anthony Hart, 22-Oct-2010.)
(( ⊥ → ⊥ ) ↔ ⊤ )
 
Theoremnottru 1348 A ¬ identity. (Contributed by Anthony Hart, 22-Oct-2010.)
(¬ ⊤ ↔ ⊥ )
 
Theoremnotfal 1349 A ¬ identity. (Contributed by Anthony Hart, 22-Oct-2010.) (Proof shortened by Andrew Salmon, 13-May-2011.)
(¬ ⊥ ↔ ⊤ )
 
Theoremtrubitru 1350 A identity. (Contributed by Anthony Hart, 22-Oct-2010.) (Proof shortened by Andrew Salmon, 13-May-2011.)
(( ⊤ ↔ ⊤ ) ↔ ⊤ )
 
Theoremtrubifal 1351 A identity. (Contributed by Anthony Hart, 22-Oct-2010.) (Proof shortened by Andrew Salmon, 13-May-2011.)
(( ⊤ ↔ ⊥ ) ↔ ⊥ )
 
Theoremfalbitru 1352 A identity. (Contributed by Anthony Hart, 22-Oct-2010.) (Proof shortened by Andrew Salmon, 13-May-2011.)
(( ⊥ ↔ ⊤ ) ↔ ⊥ )
 
Theoremfalbifal 1353 A identity. (Contributed by Anthony Hart, 22-Oct-2010.) (Proof shortened by Andrew Salmon, 13-May-2011.)
(( ⊥ ↔ ⊥ ) ↔ ⊤ )
 
Theoremtrunantru 1354 A identity. (Contributed by Anthony Hart, 22-Oct-2010.) (Proof shortened by Andrew Salmon, 13-May-2011.)
(( ⊤ ⊤ ) ↔ ⊥ )
 
Theoremtrunanfal 1355 A identity. (Contributed by Anthony Hart, 23-Oct-2010.) (Proof shortened by Andrew Salmon, 13-May-2011.)
(( ⊤ ⊥ ) ↔ ⊤ )
 
Theoremfalnantru 1356 A identity. (Contributed by Anthony Hart, 23-Oct-2010.) (Proof shortened by Andrew Salmon, 13-May-2011.)
(( ⊥ ⊤ ) ↔ ⊤ )
 
Theoremfalnanfal 1357 A identity. (Contributed by Anthony Hart, 22-Oct-2010.) (Proof shortened by Andrew Salmon, 13-May-2011.)
(( ⊥ ⊥ ) ↔ ⊤ )
 
Theoremtruxortru 1358 A identity. (Contributed by David A. Wheeler, 8-May-2015.)
(( ⊤ ⊻ ⊤ ) ↔ ⊥ )
 
Theoremtruxorfal 1359 A identity. (Contributed by David A. Wheeler, 8-May-2015.)
(( ⊤ ⊻ ⊥ ) ↔ ⊤ )
 
Theoremfalxortru 1360 A identity. (Contributed by David A. Wheeler, 9-May-2015.)
(( ⊥ ⊻ ⊤ ) ↔ ⊤ )
 
Theoremfalxorfal 1361 A identity. (Contributed by David A. Wheeler, 9-May-2015.)
(( ⊥ ⊻ ⊥ ) ↔ ⊥ )
 
1.2.13  Auxiliary theorems for Alan Sare's virtual deduction tool, part 1
 
Theoremee22 1362 Virtual deduction rule e22 in set.mm without virtual deduction connectives. Special theorem needed for Alan Sare's virtual deduction translation tool. (Contributed by Alan Sare, 2-May-2011.) (New usage is discouraged.) TODO: decide if this is worth keeping.
(φ → (ψχ))    &   (φ → (ψθ))    &   (χ → (θτ))       (φ → (ψτ))
 
Theoremee12an 1363 e12an in set.mm without virtual deduction connectives. Special theorem needed for Alan Sare's virtual deduction translation tool. (Contributed by Alan Sare, 28-Oct-2011.) TODO: this is frequently used; come up with better label.
(φψ)    &   (φ → (χθ))    &   ((ψ θ) → τ)       (φ → (χτ))
 
Theoremee23 1364 e23 in set.mm without virtual deductions. (Contributed by Alan Sare, 17-Jul-2011.) (New usage is discouraged.) TODO: decide if this is worth keeping.
(φ → (ψχ))    &   (φ → (ψ → (θτ)))    &   (χ → (τη))       (φ → (ψ → (θη)))
 
Theoremexbir 1365 Exportation implication also converting head from biconditional to conditional. This proof is exbirVD in set.mm automatically translated and minimized. (Contributed by Alan Sare, 31-Dec-2011.) (New usage is discouraged.) TODO: decide if this is worth keeping.
(((φ ψ) → (χθ)) → (φ → (ψ → (θχ))))
 
Theorem3impexp 1366 impexp 433 with a 3-conjunct antecedent. (Contributed by Alan Sare, 31-Dec-2011.)
(((φ ψ χ) → θ) ↔ (φ → (ψ → (χθ))))
 
Theorem3impexpbicom 1367 3impexp 1366 with biconditional consequent of antecedent that is commuted in consequent. Derived automatically from 3impexpVD in set.mm. (Contributed by Alan Sare, 31-Dec-2011.) (New usage is discouraged.) TODO: decide if this is worth keeping.
(((φ ψ χ) → (θτ)) ↔ (φ → (ψ → (χ → (τθ)))))
 
Theorem3impexpbicomi 1368 Deduction form of 3impexpbicom 1367. Derived automatically from 3impexpbicomiVD in set.mm. (Contributed by Alan Sare, 31-Dec-2011.) (New usage is discouraged.) TODO: decide if this is worth keeping.
((φ ψ χ) → (θτ))       (φ → (ψ → (χ → (τθ))))
 
Theoremancomsimp 1369 Closed form of ancoms 439. Derived automatically from ancomsimpVD in set.mm. (Contributed by Alan Sare, 31-Dec-2011.)
(((φ ψ) → χ) ↔ ((ψ φ) → χ))
 
Theoremexp3acom3r 1370 Export and commute antecedents. (Contributed by Alan Sare, 18-Mar-2012.)
(φ → ((ψ χ) → θ))       (ψ → (χ → (φθ)))
 
Theoremexp3acom23g 1371 Implication form of exp3acom23 1372. (Contributed by Alan Sare, 22-Jul-2012.) (New usage is discouraged.) TODO: decide if this is worth keeping.
((φ → ((ψ χ) → θ)) ↔ (φ → (χ → (ψθ))))
 
Theoremexp3acom23 1372 The exportation deduction exp3a 425 with commutation of the conjoined wwfs. (Contributed by Alan Sare, 22-Jul-2012.)
(φ → ((ψ χ) → θ))       (φ → (χ → (ψθ)))
 
Theoremsimplbi2comg 1373 Implication form of simplbi2com 1374. (Contributed by Alan Sare, 22-Jul-2012.) (New usage is discouraged.) TODO: decide if this is worth keeping.
((φ ↔ (ψ χ)) → (χ → (ψφ)))
 
Theoremsimplbi2com 1374 A deduction eliminating a conjunct, similar to simplbi2 608. (Contributed by Alan Sare, 22-Jul-2012.) (Proof shortened by Wolf Lammen, 10-Nov-2012.)
(φ ↔ (ψ χ))       (χ → (ψφ))
 
Theoremee21 1375 e21 in set.mm without virtual deductions. (Contributed by Alan Sare, 18-Mar-2012.) (New usage is discouraged.) TODO: decide if this is worth keeping.
(φ → (ψχ))    &   (φθ)    &   (χ → (θτ))       (φ → (ψτ))
 
Theoremee10 1376 e10 in set.mm without virtual deductions. (Contributed by Alan Sare, 25-Jul-2011.) TODO: this is frequently used; come up with better label.
(φψ)    &   χ    &   (ψ → (χθ))       (φθ)
 
Theoremee02 1377 e02 in set.mm without virtual deductions. (Contributed by Alan Sare, 22-Jul-2012.) (New usage is discouraged.) TODO: decide if this is worth keeping.
φ    &   (ψ → (χθ))    &   (φ → (θτ))       (ψ → (χτ))
 
1.2.14  Half-adders and full adders in propositional calculus

Propositional calculus deals with truth values, which can be interpreted as bits. Using this, we can define the half-adder in pure propositional calculus, and show its basic properties.

 
Syntaxwhad 1378 Define the half adder (triple XOR). (Contributed by Mario Carneiro, 4-Sep-2016.)
wff hadd(φ, ψ, χ)
 
Syntaxwcad 1379 Define the half adder carry. (Contributed by Mario Carneiro, 4-Sep-2016.)
wff cadd(φ, ψ, χ)
 
Definitiondf-had 1380 Define the half adder (triple XOR). (Contributed by Mario Carneiro, 4-Sep-2016.)
(hadd(φ, ψ, χ) ↔ ((φψ) ⊻ χ))
 
Definitiondf-cad 1381 Define the half adder carry, which is true when at least two arguments are true. (Contributed by Mario Carneiro, 4-Sep-2016.)
(cadd(φ, ψ, χ) ↔ ((φ ψ) (χ (φψ))))
 
Theoremhadbi123d 1382 Equality theorem for half adder. (Contributed by Mario Carneiro, 4-Sep-2016.)
(φ → (ψχ))    &   (φ → (θτ))    &   (φ → (ηζ))       (φ → (hadd(ψ, θ, η) ↔ hadd(χ, τ, ζ)))
 
Theoremcadbi123d 1383 Equality theorem for adder carry. (Contributed by Mario Carneiro, 4-Sep-2016.)
(φ → (ψχ))    &   (φ → (θτ))    &   (φ → (ηζ))       (φ → (cadd(ψ, θ, η) ↔ cadd(χ, τ, ζ)))
 
Theoremhadbi123i 1384 Equality theorem for half adder. (Contributed by Mario Carneiro, 4-Sep-2016.)
(φψ)    &   (χθ)    &   (τη)       (hadd(φ, χ, τ) ↔ hadd(ψ, θ, η))
 
Theoremcadbi123i 1385 Equality theorem for adder carry. (Contributed by Mario Carneiro, 4-Sep-2016.)
(φψ)    &   (χθ)    &   (τη)       (cadd(φ, χ, τ) ↔ cadd(ψ, θ, η))
 
Theoremhadass 1386 Associative law for triple XOR. (Contributed by Mario Carneiro, 4-Sep-2016.)
(hadd(φ, ψ, χ) ↔ (φ ⊻ (ψχ)))
 
Theoremhadbi 1387 The half adder is the same as the triple biconditional. (Contributed by Mario Carneiro, 4-Sep-2016.)
(hadd(φ, ψ, χ) ↔ ((φψ) ↔ χ))
 
Theoremhadcoma 1388 Commutative law for triple XOR. (Contributed by Mario Carneiro, 4-Sep-2016.)
(hadd(φ, ψ, χ) ↔ hadd(ψ, φ, χ))
 
Theoremhadcomb 1389 Commutative law for triple XOR. (Contributed by Mario Carneiro, 4-Sep-2016.)
(hadd(φ, ψ, χ) ↔ hadd(φ, χ, ψ))
 
Theoremhadrot 1390 Rotation law for triple XOR. (Contributed by Mario Carneiro, 4-Sep-2016.)
(hadd(φ, ψ, χ) ↔ hadd(ψ, χ, φ))
 
Theoremcador 1391 Write the adder carry in disjunctive normal form. (Contributed by Mario Carneiro, 4-Sep-2016.)
(cadd(φ, ψ, χ) ↔ ((φ ψ) (φ χ) (ψ χ)))
 
Theoremcadan 1392 Write the adder carry in conjunctive normal form. (Contributed by Mario Carneiro, 4-Sep-2016.)
(cadd(φ, ψ, χ) ↔ ((φ ψ) (φ χ) (ψ χ)))
 
Theoremhadnot 1393 The half adder distributes over negation. (Contributed by Mario Carneiro, 4-Sep-2016.)
(¬ hadd(φ, ψ, χ) ↔ hadd(¬ φ, ¬ ψ, ¬ χ))
 
Theoremcadnot 1394 The adder carry distributes over negation. (Contributed by Mario Carneiro, 4-Sep-2016.)
(¬ cadd(φ, ψ, χ) ↔ cadd(¬ φ, ¬ ψ, ¬ χ))
 
Theoremcadcoma 1395 Commutative law for adder carry. (Contributed by Mario Carneiro, 4-Sep-2016.)
(cadd(φ, ψ, χ) ↔ cadd(ψ, φ, χ))
 
Theoremcadcomb 1396 Commutative law for adder carry. (Contributed by Mario Carneiro, 4-Sep-2016.)
(cadd(φ, ψ, χ) ↔ cadd(φ, χ, ψ))
 
Theoremcadrot 1397 Rotation law for adder carry. (Contributed by Mario Carneiro, 4-Sep-2016.)
(cadd(φ, ψ, χ) ↔ cadd(ψ, χ, φ))
 
Theoremcad1 1398 If one parameter is true, the adder carry is true exactly when at least one of the other parameters is true. (Contributed by Mario Carneiro, 8-Sep-2016.)
(χ → (cadd(φ, ψ, χ) ↔ (φ ψ)))
 
Theoremcad11 1399 If two parameters are true, the adder carry is true. (Contributed by Mario Carneiro, 4-Sep-2016.)
((φ ψ) → cadd(φ, ψ, χ))
 
Theoremcad0 1400 If one parameter is false, the adder carry is true exactly when both of the other two parameters are true. (Contributed by Mario Carneiro, 8-Sep-2016.)
χ → (cadd(φ, ψ, χ) ↔ (φ ψ)))
    < Previous  Next >

Page List
Jump to page: Contents  1 1-100 2 101-200 3 201-300 4 301-400 5 401-500 6 501-600 7 601-700 8 701-800 9 801-900 10 901-1000 11 1001-1100 12 1101-1200 13 1201-1300 14 1301-1400 15 1401-1500 16 1501-1600 17 1601-1700 18 1701-1800 19 1801-1900 20 1901-2000 21 2001-2100 22 2101-2200 23 2201-2300 24 2301-2400 25 2401-2500 26 2501-2600 27 2601-2700 28 2701-2800 29 2801-2900 30 2901-3000 31 3001-3100 32 3101-3200 33 3201-3300 34 3301-3400 35 3401-3500 36 3501-3600 37 3601-3700 38 3701-3800 39 3801-3900 40 3901-4000 41 4001-4100 42 4101-4200 43 4201-4300 44 4301-4400 45 4401-4500 46 4501-4600 47 4601-4700 48 4701-4800 49 4801-4900 50 4901-5000 51 5001-5100 52 5101-5200 53 5201-5300 54 5301-5400 55 5401-5500 56 5501-5600 57 5601-5700 58 5701-5800 59 5801-5900 60 5901-6000 61 6001-6100 62 6101-6200 63 6201-6300 64 6301-6338
  Copyright terms: Public domain < Previous  Next >