section ‹Metamath introduction›
theory MMI_prelude imports Order_ZF_1
begin
text‹Metamath's set.mm features a large (over 8000) collection of theorems
proven in the ZFC set theory. This theory is part of an attempt to translate
those theorems to Isar so that they are available for Isabelle/ZF users.
A total of about 1200 assertions have been translated, 600 of that
with proofs (the rest was proven automatically by Isabelle).
The translation was done with the support of the mmisar tool, whose source
is included in the IsarMathLib distributions prior to version 1.6.4.
The translation tool was doing about 99 percent of work involved, with the rest
mostly related to the difference between Isabelle/ZF and Metamath
metalogics. Metamath uses Tarski-Megill metalogic that does not have a notion
of bound variables (see
‹ http://planetx.cc.vt.edu/AsteroidMeta/Distinctors_vs_binders ›
for details and discussion).
The translation project is closed now as I decided that it was too boring
and tedious even with the support of mmisar software. Also, the translated
proofs are not as readable as native Isar proofs which goes against IsarMathLib
philosophy.›
subsection‹Importing from Metamath - how is it done›
text‹
We are interested in importing the theorems about complex numbers
that start from the "recnt" theorem on. This is done mostly automatically
by the mmisar tool that is included in the IsarMathLib distributions prior
to version 1.6.4.
The tool works as follows:
First it reads the list of (Metamath) names of theorems
that are already imported to IsarMathlib ("known theorems")
and the list of theorems that are intended to be imported in this
session ("new theorems").
The new theorems are consecutive theorems about complex numbers
as they appear in the Metamath database.
Then mmisar creates a "Metamath script" that contains
Metamath commands that open a log file and put the statements
and proofs of the new theorems in that file in a readable format.
The tool writes this script to a disk file and executes metamath
with standard input redirected from that file. Then the log file is read
and its contents converted to the Isar format. In Metamath,
the proofs of theorems about complex numbers
depend only on 28 axioms of complex numbers and some basic logic and
set theory theorems.
The tool finds which of these dependencies are not known yet and repeats
the process of getting their statements from Metamath as with the
new theorems. As a result of this process mmisar creates files
new\_theorems.thy, new\_deps.thy and new\_known\_theorems.txt.
The file new\_theorems.thy contains the theorems (with proofs)
imported from Metamath in this session. These theorems are added
(by hand) to the current ‹MMI_Complex_ZF_x.thy› file.
The file new\_deps.thy contains the
statements of new dependencies with generic proofs "by auto".
These are added to the ‹MMI_logic_and_sets.thy›.
Most of the dependencies can be proven automatically by Isabelle.
However, some manual work has to be done for the dependencies
that Isabelle can not prove by itself and to correct problems related
to the fact that Metamath uses a metalogic based on
distinct variable constraints (Tarski-Megill metalogic),
rather than an explicit notion of free and bound variables.
The old list of known theorems is replaced by the new list and
mmisar is ready to convert the next batch of new theorems.
Of course this rarely works in practice without tweaking the mmisar
source files every time a new batch is processed.›
subsection‹The context for Metamath theorems›
text‹We list the Metamth's axioms of complex numbers
and define notation here.›
text‹The next definition is what Metamath $X\in V$ is
translated to. I am not sure why it works, probably because
Isabelle does a type inference and the "=" sign
indicates that both sides are sets.›
definition
IsASet :: "i⇒o" ("_ isASet" [90] 90) where
IsASet_def[simp]: "X isASet ≡ X = X"
text‹The next locale sets up the context to which Metamath theorems
about complex numbers are imported. It assumes the axioms
of complex numbers and defines the notation used for complex numbers.
One of the problems with importing theorems from Metamath is that
Metamath allows direct infix notation for binary operations so
that the notation $a f b$ is allowed where $f$ is a function
(that is, a set of pairs). To my knowledge,
Isar allows only notation ‹f`⟨a,b⟩› with a possibility of
defining a syntax say ‹a \<ca> b› to mean the same as ‹f`⟨a,b⟩›
(please correct me if I am wrong here). This is why we have
two objects for addition: one called ‹caddset› that represents
the binary function, and the second one called ‹ca› which
defines the ‹a \<ca> b› notation for ‹caddset`⟨a,b⟩›.
The same applies to multiplication of real numbers.
Another difficulty is that Metamath allows to define sets with syntax
$\{ x | p\}$ where $p$ is some formula that (usually) depends on $x$.
Isabelle allows the set comprehension like this only as a subset of another
set i.e. $\{x\in A . p(x)\}$. This forces us to have a sligtly different
definition of (complex) natural numbers, requiring explicitly that natural
numbers is a subset of reals. Because of that, the proofs of Metamath theorems
that reference the definition directly can not be imported.
›
locale MMIsar0 =
fixes real ("ℝ")
fixes complex ("ℂ")
fixes one ("𝟭")
fixes zero ("𝟬")
fixes iunit ("𝗂")
fixes caddset ("\<caddset>")
fixes cmulset ("\<cmulset>")
fixes lessrrel ("\<lsrset>")
fixes ca (infixl "\<ca>" 69)
defines ca_def: "a \<ca> b ≡ \<caddset>`⟨a,b⟩"
fixes cm (infixl "⋅" 71)
defines cm_def: "a ⋅ b ≡ \<cmulset>`⟨a,b⟩"
fixes sub (infixl "\<cs>" 69)
defines sub_def: "a \<cs> b ≡ ⋃ { x ∈ ℂ. b \<ca> x = a }"
fixes cneg ("\<cn>_" 95)
defines cneg_def: "\<cn> a ≡ 𝟬 \<cs> a"
fixes cdiv (infixl "\<cdiv>" 70)
defines cdiv_def: "a \<cdiv> b ≡ ⋃ { x ∈ ℂ. b ⋅ x = a }"
fixes cpnf ("\<cpnf>")
defines cpnf_def: "\<cpnf> ≡ ℂ"
fixes cmnf ("\<cmnf>")
defines cmnf_def: "\<cmnf> ≡ {ℂ}"
fixes cxr ("ℝ⇧*")
defines cxr_def: "ℝ⇧* ≡ ℝ ∪ {\<cpnf>,\<cmnf>}"
fixes cxn ("ℕ")
defines cxn_def: "ℕ ≡ ⋂ {N ∈ Pow(ℝ). 𝟭 ∈ N ∧ (∀n. n∈N ⟶ n\<ca>𝟭 ∈ N)}"
fixes lessr (infix "\<lsr>" 68)
defines lessr_def: "a \<lsr> b ≡ ⟨a,b⟩ ∈ \<lsrset>"
fixes cltrrset ("\<cltrrset>")
defines cltrrset_def:
"\<cltrrset> ≡ (\<lsrset> ∩ ℝ×ℝ) ∪ {⟨\<cmnf>,\<cpnf>⟩} ∪
(ℝ×{\<cpnf>}) ∪ ({\<cmnf>}×ℝ )"
fixes cltrr (infix "\<ls>" 68)
defines cltrr_def: "a \<ls> b ≡ ⟨a,b⟩ ∈ \<cltrrset>"
fixes convcltrr (infix ">" 68)
defines convcltrr_def: "a > b ≡ ⟨a,b⟩ ∈ converse(\<cltrrset>)"
fixes lsq (infix "\<lsq>" 68)
defines lsq_def: "a \<lsq> b ≡ ¬ (b \<ls> a)"
fixes two ("𝟮")
defines two_def: "𝟮 ≡ 𝟭\<ca>𝟭"
fixes three ("𝟯")
defines three_def: "𝟯 ≡ 𝟮\<ca>𝟭"
fixes four ("𝟰")
defines four_def: "𝟰 ≡ 𝟯\<ca>𝟭"
fixes five ("𝟱")
defines five_def: "𝟱 ≡ 𝟰\<ca>𝟭"
fixes six ("𝟲")
defines six_def: "𝟲 ≡ 𝟱\<ca>𝟭"
fixes seven ("𝟳")
defines seven_def: "𝟳 ≡ 𝟲\<ca>𝟭"
fixes eight ("𝟴")
defines eight_def: "𝟴 ≡ 𝟳\<ca>𝟭"
fixes nine ("𝟵")
defines nine_def: "𝟵 ≡ 𝟴\<ca>𝟭"
assumes MMI_pre_axlttri:
"A ∈ ℝ ∧ B ∈ ℝ ⟶ (A \<lsr> B ⟷ ¬(A=B ∨ B \<lsr> A))"
assumes MMI_pre_axlttrn:
"A ∈ ℝ ∧ B ∈ ℝ ∧ C ∈ ℝ ⟶ ((A \<lsr> B ∧ B \<lsr> C) ⟶ A \<lsr> C)"
assumes MMI_pre_axltadd:
"A ∈ ℝ ∧ B ∈ ℝ ∧ C ∈ ℝ ⟶ (A \<lsr> B ⟶ C\<ca>A \<lsr> C\<ca>B)"
assumes MMI_pre_axmulgt0:
"A ∈ ℝ ∧ B ∈ ℝ ⟶ ( 𝟬 \<lsr> A ∧ 𝟬 \<lsr> B ⟶ 𝟬 \<lsr> A⋅B)"
assumes MMI_pre_axsup:
"A ⊆ ℝ ∧ A ≠ 0 ∧ (∃x∈ℝ. ∀y∈A. y \<lsr> x) ⟶
(∃x∈ℝ. (∀y∈A. ¬(x \<lsr> y)) ∧ (∀y∈ℝ. (y \<lsr> x ⟶ (∃z∈A. y \<lsr> z))))"
assumes MMI_axresscn: "ℝ ⊆ ℂ"
assumes MMI_ax1ne0: "𝟭 ≠ 𝟬"
assumes MMI_axcnex: "ℂ isASet"
assumes MMI_axaddopr: "\<caddset> : ( ℂ × ℂ ) → ℂ"
assumes MMI_axmulopr: "\<cmulset> : ( ℂ × ℂ ) → ℂ"
assumes MMI_axmulcom: "A ∈ ℂ ∧ B ∈ ℂ ⟶ A ⋅ B = B ⋅ A"
assumes MMI_axaddcl: "A ∈ ℂ ∧ B ∈ ℂ ⟶ A \<ca> B ∈ ℂ"
assumes MMI_axmulcl: "A ∈ ℂ ∧ B ∈ ℂ ⟶ A ⋅ B ∈ ℂ"
assumes MMI_axdistr:
"A ∈ ℂ ∧ B ∈ ℂ ∧ C ∈ ℂ ⟶ A⋅(B \<ca> C) = A⋅B \<ca> A⋅C"
assumes MMI_axaddcom: "A ∈ ℂ ∧ B ∈ ℂ ⟶ A \<ca> B = B \<ca> A"
assumes MMI_axaddass:
"A ∈ ℂ ∧ B ∈ ℂ ∧ C ∈ ℂ ⟶ A \<ca> B \<ca> C = A \<ca> (B \<ca> C)"
assumes MMI_axmulass:
"A ∈ ℂ ∧ B ∈ ℂ ∧ C ∈ ℂ ⟶ A ⋅ B ⋅ C = A ⋅ (B ⋅ C)"
assumes MMI_ax1re: "𝟭 ∈ ℝ"
assumes MMI_axi2m1: "𝗂 ⋅ 𝗂 \<ca> 𝟭 = 𝟬"
assumes MMI_ax0id: "A ∈ ℂ ⟶ A \<ca> 𝟬 = A"
assumes MMI_axicn: "𝗂 ∈ ℂ"
assumes MMI_axnegex: "A ∈ ℂ ⟶ ( ∃ x ∈ ℂ. ( A \<ca> x ) = 𝟬 )"
assumes MMI_axrecex: "A ∈ ℂ ∧ A ≠ 𝟬 ⟶ ( ∃ x ∈ ℂ. A ⋅ x = 𝟭)"
assumes MMI_ax1id: "A ∈ ℂ ⟶ A ⋅ 𝟭 = A"
assumes MMI_axaddrcl: "A ∈ ℝ ∧ B ∈ ℝ ⟶ A \<ca> B ∈ ℝ"
assumes MMI_axmulrcl: "A ∈ ℝ ∧ B ∈ ℝ ⟶ A ⋅ B ∈ ℝ"
assumes MMI_axrnegex: "A ∈ ℝ ⟶ ( ∃ x ∈ ℝ. A \<ca> x = 𝟬 )"
assumes MMI_axrrecex: "A ∈ ℝ ∧ A ≠ 𝟬 ⟶ ( ∃ x ∈ ℝ. A ⋅ x = 𝟭 )"
end