SlideShare ist ein Scribd-Unternehmen logo
1 von 70
Downloaden Sie, um offline zu lesen
Mizarによる形式証明の簡単 な実例 (といろいろ宣伝) 
Stowarzyszenie Uzytkownikow Mizara (信州大学) 
↑ポーランド語でASSOCIATION OF MIZAR USERSらしい 
@Alcor80UMa (岡崎裕之)
Mizar 大熊座(Ura Major略してUma)ζ 星:79番星 なお、Alcorは80番星
目的:Mizarの宣伝 
•ProofSummit 2014 
定理証明支援系 
(Coq,Agda,Isabelle/HOL,ACL2などなど) 
↑ここ 
や自動証明器などを中心に集まるお祭りです。 
2014/9/7
Formalizing 100 Theorems http://www.cs.ru.nl/~freek/100/ 
1.HOL Light 86 
2.Mizar 61 
3.Isabelle 54 
4.Coq 49 
5.以下略 
全体での達成率89%
やっぱりアウェーな感じですが気 にせずやってみましょう 
お願い: 
単なるユーザーですので難しい話は聞かない でください。 
(暗号の人なのでこのあたりに居ます。 
日本応用数理学会 
「数理的技法による情報セキュリティ」研究部会 
http://nicosia.is.s.u-tokyo.ac.jp/jsiam-fais/ 
) 
2014/9/7
begin 中正さんのスライド: Mizarプロジェクト概要と現状 
日本における参加状況 
(スライド提供:信州大学D2中正和久)
形式化数学とは 
•概念 
–理論とは、「記号列を操作するゲーム」である 
–公理とは、記号列を操作するルールである 
–定理とは、ルールに則って得られる記号列である 
–証明とは、定理を得るまでの記号列の操作手順である 
•歴史 
–古典的な形式化数学 
•ヒルベルト・プログラム(1920) 
•ブルバキ「数学原論」(1935-) 
–計算機支援による形式化数学(1960s-) 
•Mizar(1973)
近年の背景事情 
1.現代数学の高度化・複雑化 
–Hilbert~Grothendieckを流れとした数学の抽象化と高度な発展 
–数論幾何・超弦理論など高度に発展した分野同士の融合 
2.システムの堅牢性への要求の高まり 
–暗号理論・セキュリティなど、現代数学がシステムの堅牢性を支えている 
–数学自体の堅牢性を確かめておかなければならない 
3.情報科学の成熟と新たな成功 
–クラウドコンピューティング(Google) 
–集合知(Wikipedia) 
•現代数学の知識体系を計算機で処理可能なコーパスにしようという機運 は確実に高まっている(MKM: Mathematical Knowledge Management) 
–QED manifesto (1994) 
–MKM-IG (2001)
Mizarプロジェクト 
1.Mizar 
言語 
3.MML 
2.Mizar 
Verifier 
Mizarで書かれた証明を 
検証するプログラム 
数学の証明を 
記述する言語 
証明済みの定義や定理を 
集めたライブラリ 
特長1. 
高い可読性 
特長3. 
膨大な資産 
特長2. 
検証に特化 
閲覧・引用 
検証 
投稿
Mizarの特長1. 高い可読性 
•数学証明記述に特化 
•自然言語に近い文法 
•一階述語論理 
•宣言型言語 
•柔軟な型定義 
•記号が少ない 
Freek Wiedijk: Comparing mathematical provers より引用
Mizarの特長2. 検証に特化 
•他の多くのシステム(HOL, Coq, …)は自動証明に力 点を置いている 
•Mizarは証明のチェックに力点を置いている 
–数学知識体系のライブラリ化が主目的 
–人間と計算機にそれぞれ得意な分野を分担させる思想。 証明の記述は人間が、検証は計算機が行う。 
ませまてぃしゃんの人が趣味で作ったのでこうなった? 
集合理論ベース(選択公理)で一階述語論
Mizarの特長3. 膨大な資産 
•論文誌Formalized Mathematicsは1990年発刊 
•11,600の定義と53,700の定理、261万行 (2013年 7月時点) 
•オープンライセンス(GNU v3/CC BY-SA v3)でWeb 上に公開 
–http://mizar.org/library/ 
•現在も活発に執筆が続けられている
Mizarの記述例(TAYLOR展開定理) 
theorem :: TAYLOR_1:33 
for n be Nat, f be PartFunc of REAL,REAL, x0,r be Real st 
( 0 < r & f is_differentiable_on n+1, ].x0-r,x0+ r.[ ) 
for x be Real st x in ].x0-r, 
x0+r.[ holds ex s be Real st 0 < s & s< 1 & 
f.x=Partial_Sums(Taylor(f, ].x0-r,x0+r.[,x0,x)).n + (diff(f,].x0- r,x0+r.[).(n+1)).(x0+s*(x-x0)) * (x-x0) |^ (n+1) / ((n+1)!);
他システムで記述すると 
|- !f diff h n. &0 < h /¥ 0 < n /¥ (diff(0) = f) /¥ (!m t. m < n /¥ &0 <= t /¥ t <= h ==> (diff(m) diffl diff(SUC m)(t))(t)) ==> (?t. &0 < t /¥ t < h /¥ (f(h) = sum(0,n)(¥m. (diff(m)(&0) / &(FACT m)) * (h pow m)) + ((diff(n)(t) / &(FACT n)) * (h pow n)))) 
判読が難しい 
(例 HOL Light)
膨大な数学定理の形式証明ライブラリ を作るには 
•どの定理や定義をどのように並べ,有機的に関 連づけていくのか? 
•どのように形式証明を記述していくか? 
•組織立てられた作成計画が必須 
•デザインなど人間の高次な脳の機能に依存 
•視野が観ている窓に限定される対話型の定理 証明支援系では困難
現代数学の大規模なライブラリ構築 に重点をおくなら 
現状の計算機の技術では, 
•形式証明自体は人間が書く 
•計算機はそれを精査する 
–それぞれ得意なものに集中する役割分業が有効 
•QEDの目的達成にはほとんどの数理・情報科学の 研究者が,自分が執筆する通常の論文と同様に, 読み書きに不自由しないシステムの開発が重要
MML(mizar mathematical library) music macro languageではない 
2013年6月時点で, 
Mizar Mathematical Library(MML)は 
•延べ200名を超える著者 
•1181のファイル 
•約52,000の定理 
•10,000を超える形式上の定義 
WEB上に公開 
http://mizar.uwb.edu.pl/version/current/html/
BiaLystok大学(ポーランド), Alberta大学(カナダ), 
信州大学(日本)で研究グループによって展開, 
推進されている 
http://mizar.org/project/ 
チュートリアル, ライブラリ検索, 各種リソース 
markun.cs.shinshu-u.ac.jp/kiso/projects/proofchecker/mizar/ 
index-j.html 
成果と資源はweb上ですべて公開
形式化記述の例(環境設定部分) 
environ Vocabularies NUMBERS, INT_1, ORDINAL1, SUBSET_1, CARD_1, ARYTM_3, ARYTM_1, RELAT_1, NAT_1,XCMPLX_0,REAL_1; Notations SUBSET_1,ORDINAL1,NUMBERS, XCMPLX_0, INT_1, NAT_1; constructors REAL_1,NAT_1, INT_1, CARD_1; Registrations XREAL_0,NAT_1, INT_1, ORDINAL1, XBOOLE_0; Requirements REAL, NUMERALS, SUBSET,ARITHM, BOOLE; definitions INT_1; Theorems INT_1; Schemes NAT_1;
環境部には 
•Vocabularies 用語の定義ファイルを指定する。 
•Notations 
•constructors 個々の概念の定義する 概念間の階層構造等の関係を表す。 
•registrations 変数型の階層的な関係,数学的属性の従属関係などを指定する。 
•Requirements 実数の演算や集合の包含関係の演算の自動的な処理を指定する。 
•definitions 
•theorems 
•schemes begin文以下の形式化記述の本体で引用される既存の定義や定理, 公理図式の収録ファイルを指定する。
形式化記述の例(本体部分) 
begin 
theorem EX1: 
for n being Nat holds 2 divides n*(n-1) 
proof 
defpred P[Nat] means 2 divides $1*($1-1); 
0 = 2 * 0; 
then 
P0: P[0] by INT_1:def 3; 
PN: for n being Nat st P[n] holds P[n+1] 
proof 
let n be Nat; 
assume P[n]; 
then 
consider s being Integer such that 
A2: n * (n - 1) = 2 * s by INT_1:def 3 
(n+1)*(n+1-1) = n*(n-1) + n + n*1 
.=n*(n-1) + 2*n 
.=2*s + 2*n by A2 
.= 2*(s+n); 
hence P[n+1] by INT_1:def 3; 
end; 
thus for n being Nat holds P[n] from NAT_1:sch 2(P0,PN); 
end;
theorem EX1: 
命題名EX1という命題のラベル名の宣 言 
for n being Nat holds 2 divides n*(n-1) 
「任意の自然数nに対して,n・(n-1)は 2で割り切れる。」という命題 
(注) Natは自然数の変数型名であ る。 
proof 
命題EX1の証明開始 
(注) 以下proofとend;で挟まれた 
proof ~ endブロックが書かれ,上記 命題EX1の証明が記述される。 
defpred P[Nat] means 2 divides $1*($1-1); 
自然数を引数に持つ述語P [ ] で,引 数($1)と$1-1との積($1・($1-1))が,2で 割り切れることを表わす述語を定義 する。(defpred) 
0 = 2 * 0; 
then 
P0: P[0] by INT_1:def 3; 
0 = 2・0; 
だから 2 divides 0・(0-1); であり 
補題名P0 : P[0] が 定義INT_1:def 3 により成り立っている。
PN: for n being Nat st P[n] holds P[n+1] 
補題PN: 
任意の自然数nに対して「P[n]が成り 立つときP[n+1]が成り立つ」 
proof 
補題PNの証明開始 
proof ~ endブロックで記述する。 
let n be Nat; 
nを任意の自然数とする 
assume P[n]; 
P[n] すなわち2 divides n・(n-1) 
を仮定する 
then 
consider s being Integer such that 
A2: n * (n – 1) = 2 * s by INT_1:def 3; 
従って INT_1:def 3によれば以下の補 題A2が成り立つある自然数 s を考え ることができる。 
補題A2: n・(n - 1) = 2・s
(n+1)*(n+1-1) = n*(n-1) + n + n*1 
.=n*(n-1) + 2*n 
.=2*s + 2*n by A2 
.= 2*(s+n); 
補題2などにより,等式 
(n+1)・(n+1-1) = 2・(s+n) 
が得られる 
hence 
P[n+1] by INT_1:def 3; 
よってP[n+1] がINT_1:def 3によって成 り立つ 
end; 
補題PNの証明終了 
thus for n being Nat holds 
P[n] from NAT_1:sch 2(P0,PN); 
以上補題P0,PNと数学的帰納法の公理 図式NAT_1:sch 2を用いて, 任意の 自然数nについてP[n]が成り立つ 
end; 
定理EX1の証明終了
英文に近い表現の記述を眺めれば 
Natが自然数を表すものと分かれば 
『任意の自然数nについて,2はn・(n-1)を割り 切る』という命題とそれの数学的帰納法による 証明を記述していることが大凡判読できる 
Mizarは自然言語に近い表現ができることが特徴
Mizarライブラリ(MML)の現状 
•2013/06/18現在で1181の,数学ライブラリが整備されている 
•システムとライブラリは,インターネット上で全て公開されている。 http://Mizar.uwb.edu.pl/version/current/html/ 
•数学の公理体系から出発して必要な定理に至る全ての結果・証明が完 全に収録され,しかも他を参照する必要のない自己完結的な形式化数 学記述のライブラリとして提供されている。 
•WEB表示中のproof という単語をクリックすれば,証明の全文が表示さ れる。さらにそれで引用されるライブラリ中の既存の定理や用語にもリン クが張られ,これらのリンクを繰り返し辿れば,集合論の公理系にまで戻 ることができる。
情報工学関連(日本側が始めた活動) 
DIST_* 
離散集合上の確率変数 
DESCIP_1,AESCIP_1 
秘密鍵暗号アルゴリズム 
CIRCUIT*,ftacell1, gfacirc* 
回路,並列回路 
PERTRI* 
ペトリネット 
morph_01 
画像処理と数理形態学 
NTALOG_1 
拡張ユーグリッド互除法 
ZMOD_* ECPF_* 
整数環上の Module,楕円 曲線 
AMI_*,SCM* 
抽象計算機
WEB上で利用できるMizarシステム(1) 
•主に修士の学生の講義用に使用している。 
•新しく定義するか特定のシステムコマンドを使用するかどうかによって Verifierのパッケージを選択 
•次の入力画面で使用する学生のIDやe-mailアドレスを入力する。 
•Verifierのメインページが表示される。 
証明を記入するテキストボックス 
ファイル名などを記入する入力ボックス 
がある。 
• データを入力した後SUBMITボタンをマウスでクリックすると,Mizarの proof checker が働いて, 証明をチェックし誤りがあれば,エラーメッセー ジが表示される 
•とはいってもインストールも簡単。Install.bat実行して 
環境変数2つ(MIZFILES,MIZPATH)追加するだけ。 
http://www.wakasato.jp/Mizar/
WEB上で利用できるMizarシステム(2) 
•Mizarの開発者の一人であるGrzegorz Bancerek 教授がJSPSの客員研究として本学 に招聘した際に作成 
•ライブラリの高度な検索機能 
http://mmlquery.Mizar.org/mmlquery/three.html
WEB上で利用できるMizarシステム(3) 
•Mizar Projectの重要なメンバーであるJosef Urban教授が開 発 
•証明記述中の既存のMMLの用語,引用定理へ全てリンクが 張られている 
•E-learning教材を開発する際には,他を参照する必要のない 自己完結的な形式化数学記述のライブラリとして提供できる 
•Mizarの既存のライブラリと一体で数学の公理体系から出発 して必要な定理に至る全ての結果・証明を完全に収録した教 材が作成できる 
http://Mizar.cs.ualberta.ca/~mptp/Mizar.html
終わりに 
Mizarの概要 
プロジェクトの現状と日本おける参加状況 
数理的思考訓練用の教材を汎用 CMSに組み込む研究 e-learningを主体とした社会人向け のインターネットによる遠隔教育 
暗号理論や情報工学分野での 形式検証に必要な ライブラリの整備を当面の目標に ライブラリ作成に取り組んでいる. 
学部や修士学生が学習する微積分や関数解析, 代数等の基礎的定理のライブラリ整備
•e-learning教材を開発する際に他を 参照する必要のない自己完結的な 教材を提供できる. 
•Mizarの既存のライブラリと一体で 数学の公理体系から出発して必要 な定理に至る全ての結果・ 証明 を収録した教材の作成ができる. 
数理的思考訓練用の 教材を汎用CMSに組 み込む研究 e-learningを主体とし た社会人向けのイン ターネットによる遠隔 教育 
QED manifestoの目的の一つである 
数学的な思考能力の開発に合致する. 
終わりに
見果てぬ夢 
•殆どの数学定理のライブラリ化が完成し,数学,数理科学, 応用数学分野の研究者が,形式化数学記述システムを用い て自分たちの最新の研究成果とその証明を記述できるよう になることを期待している. 
•証明の正しさは,どのような無名な研究者の論文であっても, 学会の権威者ではなく,計算機が瞬時に判定する. 
•従来の人間の査読は人間系によってのみ可能な結果の有 用性,新規性といったより高度な判定に利用される. 
•プロ,アマチュアの境がなくなり,学問研究の数理科学分野 でのユビキタスが実現されることを期待している. 
end 中正さんのスライド;
簡単な例をもう一つ 
theorem 
for x be Element of NAT st 1 < x holds 
not ex N,c be Element of NAT st 
for n be Element of NAT st N <= n holds 
x to_power n <= c * ( n to_power x) 
2014/9/7 
xnncxholdsnNtsntsNcholdsxtsx     ...., 1.. NNN
つまりこれを言いたい 
2014/9/7 
( ) n x x Ο n 
計算量の評価 
とか 
暗号の安全性の議論に使う
theorem N2POWINPOLY: 
for x be Element of NAT st 1 < x holds 
not ex N,c be Element of NAT st 
for n be Element of NAT st N <= n holds 
2 to_power n <= c * ( n to_power x) 
を使って 
for x be Element of NAT st 1 < x holds 
not ex N,c be Element of NAT st 
for n be Element of NAT st N <= n holds 
x to_power n <= c * ( n to_power x) を証明 
2014/9/7
proof 
let x be Element of NAT; 
assume AS: 1 < x; 
assume CNT: ex N,c be Element of NAT st 
for n be Element of NAT st N <= n holds 
x to_power n <= c * ( n to_power x); 
まず命題が偽であると仮定する 
2014/9/7
ex N,c be Element of NAT st 
for n be Element of NAT st N <= n holds 
2 to_power n <= c * ( n to_power x); 
(proofは次ページに) 
hence contradiction by AS,N2POWINPOLY; 
end; 
変なことが証明できたので定理N2POWINPOLY に矛盾して背理法で証明終わり 
2014/9/7
ex N,c be Element of NAT st for n be Element of NAT st N <= n holds 
2 to_power n <= c * ( n to_power x) 
proof 
consider N,c be Element of NAT such that CNT2: 
for n be Element of NAT st N <= n holds 
x to_power n <= c * ( n to_power x) by CNT; 
take N,c; 
for n be Element of NAT st N <= n holds 2 to_power n <= c * ( n to_power x) 
proof 
let n be Element of NAT; assume N <= n;then 
LCX1: x to_power n <= c * ( n to_power x) by CNT2; 
1+1 <= x by AS,INT_1:7;then 
2 to_power n <= x to_power n by LEMC01; 
hence thesis by LCX1,XXREAL_0:2; 
end; 
hence thesis; end; 
2014/9/7
数学の証明にしか使えない? 
•そういうわけでもない 
–アルゴリズム 
–回路の検証 
等にも使えます。 
例えば暗号とか 
2014/9/7
Data Encryption Standard (DES) Data Encryption Standard (DES) Symmetric cryptosystem (Block cipher). Selected by the National Bureau of Standards as an official Federal Information Processing Standard for the United States in 1976 (FIPS46). 
3 Strong influence on the design of its successors. Insecure Advanced Encryption Standard (AES) 
Now
Formal Verification of DES Using the Mizar Proof Checker 
Hiroyuki Okazaki1, Kenichi Arai2, Yasunari Shidama1 
1. Shinshu University 2. Nagano Technical High School(当時) 現:東京理科大学 
July 18, 2011
Agenda Proof is verified by using a proof checker. The correctness of our formalization of the DES algorithm. Prove the security of cryptographic systems by using the Mizar proof checker. Introduction : Our formalization of the Data Encryption Standard (DES) algorithm. 
1
Proof Checker “Mizar” The definitions and the theorems have been verified for correctness using the Mizar proof checking system. Mizar Mizar is an advanced project by Mizar Society that Prof. A. Trybulec leads in Bialystok university to which is formalized mathematics by the computer-aid. The Mizar project describes mathematical proofs in the Mizar language, which is created to formally describe mathematics. What formalizes the proof of mathematics and describes it is called article. 
2
4 
Structure of DES 
Figure 1: Structure of DES 64bits length plaintext block. 64bits length secret key. 64bits length ciphertext block. 16 rounds of processing iterations 
Decryption Encryption 
Same Key 
Feistel Structure 
64 bits 
64 bits 
64 bits
i-th Round of Feistel Structure 
5 
Figure 2: i-th round of Feistel structure 
( ) Ri-1 : 32 bits length block. Li-1 : 32 bits length block. Ki : i-th round key (48 bits). 
Exclusive OR (XOR)
About “functor” and “Function” 
6 
Mizar Two ways to define computational routines in an algorithmic sense. 
functor 
function 
A “functor” is a relation between the input and output of a routine. 
A “function” is a map from the space 
of the input onto that of the output.
Strategy of Formalizing DES in Mizar 
7 Step1 : Formalization of the algorithm of generalized DES. Step2 : Formalization of the primitives of DES according to FIPS46–3. 
Formalization of the DES algorithm
Formalization of the generalized algorithm of DES 
8 
Definition 5.1: ( Codec of gereralized DES ) 
Mizar language 
it = DES-like- CoDec(M,F,IP,RK) 
: Concatenation 
: Inverse 
[:A, B:] : Cartesian product of A and B
Correctness of the generalized algorithm of DES 
Theorem 5.1: ( Correctness of generalized DES ) 
9 
Proved that the ciphertext encoded by any Feistel cipher algorithm can be decoded uniquely with the same algorithm and secret key that were used in encryption. 
Mizar language
Formalization of DES Formalization of the DES algorithm according to FIPS46–3 in Mizar language. Step1: Formalization of the DES primitives according to FIPS46-3. Step2: Formalization of the correctness of the DES algorithm. (Using the formalization of the generalized DES algorithm.) 
Prove the correctness of the DES algorithm 
10
11 
Formalization of DES Primitives S-Boxes ( S-Box S1, ... , S-Box S5, .. , S-Box S8 ). Initial Permutation (IP). Final Permutation (IP -1). Feistel Function : ( ). Bit selection Function (E). Permutation (P). Key Scheduling Function (KS). Permuted Choice 1 (PC1). Permuted Choice 2 (PC2). Left Shift. 
DES Primitives
S-Boxes 
12 
S1 
S2 
… 
S7 
S8 
…
13 
S-Boxes 
Theorem 6.1: ( S-Box S1) 
Similarly defined the other S-Boxes, DES-SBOX2 (S2), ….. , DES-SBOX8 (S8). 
S-Box S1 
Mizar language 
it = DES-SBOX1
14 
Initial Permutation (IP) 
Definition 6.2: ( IP as functor ) 
IP 
Mizar language 
it = DES-IP(r)
Initial Permutation (IP) 
15 
Definition 6.3: ( IP as function ) 
Similarly defined the functor of the final permutation DES-IPINV and the function of the DES-PIPINV. 
Mizar language 
it = DES-PIP
16 
Feistel function 
Figure 4: Feistel function Feistel Function ( ). 
Bit selection Function (E) Permutation (P) S-Boxes 
+ 
+ 
32 bits 
48 bits 
48 bits 
48 bits 
6 bits 
6 bits 
6 bits 
6 bits 
6 bits 
6 bits 
6 bits 
6 bits 
4 bits 
4 bits 
4 bits 
4 bits 
4 bits 
4 bits 
4 bits 
4 bits 
32 bits 
32 bits
17 
Bit Selection Function (E) 
E 
Definition 6.4: ( E as functor ) 
Mizar language 
it = DES-E(r)
Permutation (P) 
18 
P 
Definition 6.5: ( P as functor ) 
Mizar language 
it = DES-P(r)
Formalization of Feistel Function 
Definition 6.6: ( Feistel function as functor ) 
19 
Mizar language 
it = DES-F(R, Rkey)
20 
Formalization of Feistel Function 
Definition 6.7: ( Feistel function as function ) 
Mizar language 
it = DES-FFUNC
21 
Key Scheduling Function (KS) 
Figure 5: Key Scheduling Function Key Scheduling Function (KS). 
Permuted Choice 1 (PC1) Permuted Choice 2 (PC2) Left Shift 
+ 
+ 
Kn = KS(n, Secret key) 
64 bits 
28 bits 
28 bits 
48 bits 
56 bits 
28 bits 
28 bits 
28 bits 
28 bits 
28 bits 
28 bits 
48 bits 
56 bits
Permuted Choice 1 (PC1) 
Definition 6.8: ( PC1 as functor ) 
Mizar language 
it = DES-PC1(r) 
PC1 
Similarly defined the functor of PC2 as DES-PC2. 
22 
C0 
D0
23 
Table of the Numbers of Left-Shift 
Definition 6.9: ( Table of Left-Shift ) 
Mizar language 
it = bitshift_DES 
i-th Numbers Round of Left-Shift 2 1 3 2 4 2 5 2 6 2 7 2 8 2 9 1 10 2 11 2 12 2 13 2 14 2 15 2 16 1 
1 1
24 
Formalization of the Key Scheduling Function 
Definition 6.7: ( Key Scheduling function ) 
Mizar language 
it = DES-KS(Key)
25 
DES Algorithm The generalized DES algorithm The DES primitives 
Definition 6.11: ( DES Algorithm ) 
Mizar language 
it = DES-CoDEC (M,F,IP,RK) 
DES algorithm
26 
Encode and Decode Algorithm of DES 
Definition 6.12: ( Encode Algorithm of DES) 
Mizar language 
Definition 6.13: ( Decode Algorithm of DES) 
Mizar language
Correctness of DES 
Theorem 6.1: ( Correctness of DES ) 
Proved using the Mizar system that the ciphertext encoded by the DES algorithm can be decoded uniquely with the same algorithm and secret key that were used in Encryption. 
27 
Mizar language
Conclusion 
28 Prove the correctness of the DES algorithm. Introduction : Our formalization of DES algorithm. 
Mizar proof checking system. 
Future work Analyze the security of DES. Prove the security of cryptographic systems by using the Mizar proof checker.
TPP告知 
TPP2014もよろしくね! 
(Theorem Proving and Provers Meeting) 
日時:2014 年 12 月 3 日(水) 〜 12 月 5 日(金) 
場所:九州大学・西新プラザ (福岡市早良区西新 2-16-23) 
講演者募集中! 
Adam Chlipala 氏 (MIT),Cyril Cohen氏 (Univ.Gothenburg) をお 招きしてCoq および Ssreflect,MathComp に関する御講演を 頂く予定だそうです。 
入口にビラが置いてありますのでよろしければお持ちください 
http://coop-math.ism.ac.jp/files/143/Announcement2.pdf 
2014/9/7

Weitere ähnliche Inhalte

Was ist angesagt?

Deep Learning を実装する
Deep Learning を実装するDeep Learning を実装する
Deep Learning を実装するShuhei Iitsuka
 
Chainer v1.6からv1.7の新機能
Chainer v1.6からv1.7の新機能Chainer v1.6からv1.7の新機能
Chainer v1.6からv1.7の新機能Ryosuke Okuta
 
「深層学習」勉強会LT資料 "Chainer使ってみた"
「深層学習」勉強会LT資料 "Chainer使ってみた"「深層学習」勉強会LT資料 "Chainer使ってみた"
「深層学習」勉強会LT資料 "Chainer使ってみた"Ken'ichi Matsui
 
PythonによるDeep Learningの実装
PythonによるDeep Learningの実装PythonによるDeep Learningの実装
PythonによるDeep Learningの実装Shinya Akiba
 
ディープニューラルネット入門
ディープニューラルネット入門ディープニューラルネット入門
ディープニューラルネット入門TanUkkii
 
2013.07.15 はじパタlt scikit-learnで始める機械学習
2013.07.15 はじパタlt scikit-learnで始める機械学習2013.07.15 はじパタlt scikit-learnで始める機械学習
2013.07.15 はじパタlt scikit-learnで始める機械学習Motoya Wakiyama
 
Clustering _ishii_2014__ch10
Clustering  _ishii_2014__ch10Clustering  _ishii_2014__ch10
Clustering _ishii_2014__ch10Kota Mori
 
Chainer/CuPy v5 and Future (Japanese)
Chainer/CuPy v5 and Future (Japanese)Chainer/CuPy v5 and Future (Japanese)
Chainer/CuPy v5 and Future (Japanese)Seiya Tokui
 
Practical recommendations for gradient-based training of deep architectures
Practical recommendations for gradient-based training of deep architecturesPractical recommendations for gradient-based training of deep architectures
Practical recommendations for gradient-based training of deep architecturesKoji Matsuda
 
Androidで動かすはじめてのDeepLearning
Androidで動かすはじめてのDeepLearningAndroidで動かすはじめてのDeepLearning
Androidで動かすはじめてのDeepLearningMiyoshi Kosuke
 
最近傍探索と直積量子化(Nearest neighbor search and Product Quantization)
最近傍探索と直積量子化(Nearest neighbor search and Product Quantization)最近傍探索と直積量子化(Nearest neighbor search and Product Quantization)
最近傍探索と直積量子化(Nearest neighbor search and Product Quantization)Nguyen Tuan
 
2015年9月18日 (GTC Japan 2015) 深層学習フレームワークChainerの導入と化合物活性予測への応用
2015年9月18日 (GTC Japan 2015) 深層学習フレームワークChainerの導入と化合物活性予測への応用 2015年9月18日 (GTC Japan 2015) 深層学習フレームワークChainerの導入と化合物活性予測への応用
2015年9月18日 (GTC Japan 2015) 深層学習フレームワークChainerの導入と化合物活性予測への応用 Kenta Oono
 
PoisoningAttackSVM (ICMLreading2012)
PoisoningAttackSVM (ICMLreading2012)PoisoningAttackSVM (ICMLreading2012)
PoisoningAttackSVM (ICMLreading2012)Hidekazu Oiwa
 
ウェーブレット木の世界
ウェーブレット木の世界ウェーブレット木の世界
ウェーブレット木の世界Preferred Networks
 
[Basic 2] 計算機の構成 / プログラム実行の仕組み
[Basic 2] 計算機の構成 / プログラム実行の仕組み[Basic 2] 計算機の構成 / プログラム実行の仕組み
[Basic 2] 計算機の構成 / プログラム実行の仕組みYuto Takei
 
[Basic 1] ブロックチェーンと計算機科学 / ブール代数 / 情報理論
[Basic 1] ブロックチェーンと計算機科学 / ブール代数 / 情報理論[Basic 1] ブロックチェーンと計算機科学 / ブール代数 / 情報理論
[Basic 1] ブロックチェーンと計算機科学 / ブール代数 / 情報理論Yuto Takei
 
Chainer, Cupy入門
Chainer, Cupy入門Chainer, Cupy入門
Chainer, Cupy入門Yuya Unno
 

Was ist angesagt? (20)

Deep Learning を実装する
Deep Learning を実装するDeep Learning を実装する
Deep Learning を実装する
 
Chainer v1.6からv1.7の新機能
Chainer v1.6からv1.7の新機能Chainer v1.6からv1.7の新機能
Chainer v1.6からv1.7の新機能
 
「深層学習」勉強会LT資料 "Chainer使ってみた"
「深層学習」勉強会LT資料 "Chainer使ってみた"「深層学習」勉強会LT資料 "Chainer使ってみた"
「深層学習」勉強会LT資料 "Chainer使ってみた"
 
NumPy闇入門
NumPy闇入門NumPy闇入門
NumPy闇入門
 
PythonによるDeep Learningの実装
PythonによるDeep Learningの実装PythonによるDeep Learningの実装
PythonによるDeep Learningの実装
 
ディープニューラルネット入門
ディープニューラルネット入門ディープニューラルネット入門
ディープニューラルネット入門
 
2013.07.15 はじパタlt scikit-learnで始める機械学習
2013.07.15 はじパタlt scikit-learnで始める機械学習2013.07.15 はじパタlt scikit-learnで始める機械学習
2013.07.15 はじパタlt scikit-learnで始める機械学習
 
Clustering _ishii_2014__ch10
Clustering  _ishii_2014__ch10Clustering  _ishii_2014__ch10
Clustering _ishii_2014__ch10
 
Chainer/CuPy v5 and Future (Japanese)
Chainer/CuPy v5 and Future (Japanese)Chainer/CuPy v5 and Future (Japanese)
Chainer/CuPy v5 and Future (Japanese)
 
Practical recommendations for gradient-based training of deep architectures
Practical recommendations for gradient-based training of deep architecturesPractical recommendations for gradient-based training of deep architectures
Practical recommendations for gradient-based training of deep architectures
 
Androidで動かすはじめてのDeepLearning
Androidで動かすはじめてのDeepLearningAndroidで動かすはじめてのDeepLearning
Androidで動かすはじめてのDeepLearning
 
最近傍探索と直積量子化(Nearest neighbor search and Product Quantization)
最近傍探索と直積量子化(Nearest neighbor search and Product Quantization)最近傍探索と直積量子化(Nearest neighbor search and Product Quantization)
最近傍探索と直積量子化(Nearest neighbor search and Product Quantization)
 
2015年9月18日 (GTC Japan 2015) 深層学習フレームワークChainerの導入と化合物活性予測への応用
2015年9月18日 (GTC Japan 2015) 深層学習フレームワークChainerの導入と化合物活性予測への応用 2015年9月18日 (GTC Japan 2015) 深層学習フレームワークChainerの導入と化合物活性予測への応用
2015年9月18日 (GTC Japan 2015) 深層学習フレームワークChainerの導入と化合物活性予測への応用
 
PoisoningAttackSVM (ICMLreading2012)
PoisoningAttackSVM (ICMLreading2012)PoisoningAttackSVM (ICMLreading2012)
PoisoningAttackSVM (ICMLreading2012)
 
ウェーブレット木の世界
ウェーブレット木の世界ウェーブレット木の世界
ウェーブレット木の世界
 
More modern gpu
More modern gpuMore modern gpu
More modern gpu
 
[Basic 2] 計算機の構成 / プログラム実行の仕組み
[Basic 2] 計算機の構成 / プログラム実行の仕組み[Basic 2] 計算機の構成 / プログラム実行の仕組み
[Basic 2] 計算機の構成 / プログラム実行の仕組み
 
[Basic 1] ブロックチェーンと計算機科学 / ブール代数 / 情報理論
[Basic 1] ブロックチェーンと計算機科学 / ブール代数 / 情報理論[Basic 1] ブロックチェーンと計算機科学 / ブール代数 / 情報理論
[Basic 1] ブロックチェーンと計算機科学 / ブール代数 / 情報理論
 
CuPy解説
CuPy解説CuPy解説
CuPy解説
 
Chainer, Cupy入門
Chainer, Cupy入門Chainer, Cupy入門
Chainer, Cupy入門
 

Ähnlich wie Proof summit2014mizar

プログラマのための文書推薦入門
プログラマのための文書推薦入門プログラマのための文書推薦入門
プログラマのための文書推薦入門y-uti
 
[DL輪読会]Factorized Variational Autoencoders for Modeling Audience Reactions to...
[DL輪読会]Factorized Variational Autoencoders for Modeling Audience Reactions to...[DL輪読会]Factorized Variational Autoencoders for Modeling Audience Reactions to...
[DL輪読会]Factorized Variational Autoencoders for Modeling Audience Reactions to...Deep Learning JP
 
深層学習フレームワーク Chainer の開発と今後の展開
深層学習フレームワーク Chainer の開発と今後の展開深層学習フレームワーク Chainer の開発と今後の展開
深層学習フレームワーク Chainer の開発と今後の展開Seiya Tokui
 
Chainerの使い方と 自然言語処理への応用
Chainerの使い方と自然言語処理への応用Chainerの使い方と自然言語処理への応用
Chainerの使い方と 自然言語処理への応用Yuya Unno
 
東京大学工学部計数工学科応用音響学 D2 Clustering
東京大学工学部計数工学科応用音響学 D2 Clustering東京大学工学部計数工学科応用音響学 D2 Clustering
東京大学工学部計数工学科応用音響学 D2 ClusteringHiroshi Ono
 
論文紹介:PaperRobot: Incremental Draft Generation of Scientific Idea
論文紹介:PaperRobot: Incremental Draft Generation of Scientific Idea論文紹介:PaperRobot: Incremental Draft Generation of Scientific Idea
論文紹介:PaperRobot: Incremental Draft Generation of Scientific IdeaHirokiKurashige
 
「TensorFlow Tutorialの数学的背景」 クイックツアー(パート1)
「TensorFlow Tutorialの数学的背景」 クイックツアー(パート1)「TensorFlow Tutorialの数学的背景」 クイックツアー(パート1)
「TensorFlow Tutorialの数学的背景」 クイックツアー(パート1)Etsuji Nakai
 
TensorFlow math ja 05 word2vec
TensorFlow math ja 05 word2vecTensorFlow math ja 05 word2vec
TensorFlow math ja 05 word2vecShin Asakawa
 
充足可能性問題のいろいろ
充足可能性問題のいろいろ充足可能性問題のいろいろ
充足可能性問題のいろいろHiroshi Yamashita
 
Rubyの御先祖CLUのお話(原本)
Rubyの御先祖CLUのお話(原本)Rubyの御先祖CLUのお話(原本)
Rubyの御先祖CLUのお話(原本)洋史 東平
 
R による文書分類入門
R による文書分類入門R による文書分類入門
R による文書分類入門Takeshi Arabiki
 
機械学習を用いたWeb上の産学連携関連文書の抽出
機械学習を用いたWeb上の産学連携関連文書の抽出機械学習を用いたWeb上の産学連携関連文書の抽出
機械学習を用いたWeb上の産学連携関連文書の抽出National Institute of Informatics
 
Hayashi tutorial ne2017
Hayashi tutorial ne2017Hayashi tutorial ne2017
Hayashi tutorial ne2017yukisachi
 
Tokyo webmining 複雑ネットワークとデータマイニング
Tokyo webmining 複雑ネットワークとデータマイニングTokyo webmining 複雑ネットワークとデータマイニング
Tokyo webmining 複雑ネットワークとデータマイニングHiroko Onari
 
Get To The Point: Summarization with Pointer-Generator Networks_acl17_論文紹介
Get To The Point: Summarization with Pointer-Generator Networks_acl17_論文紹介Get To The Point: Summarization with Pointer-Generator Networks_acl17_論文紹介
Get To The Point: Summarization with Pointer-Generator Networks_acl17_論文紹介Masayoshi Kondo
 

Ähnlich wie Proof summit2014mizar (20)

4thNLPDL
4thNLPDL4thNLPDL
4thNLPDL
 
プログラマのための文書推薦入門
プログラマのための文書推薦入門プログラマのための文書推薦入門
プログラマのための文書推薦入門
 
[DL輪読会]Factorized Variational Autoencoders for Modeling Audience Reactions to...
[DL輪読会]Factorized Variational Autoencoders for Modeling Audience Reactions to...[DL輪読会]Factorized Variational Autoencoders for Modeling Audience Reactions to...
[DL輪読会]Factorized Variational Autoencoders for Modeling Audience Reactions to...
 
深層学習フレームワーク Chainer の開発と今後の展開
深層学習フレームワーク Chainer の開発と今後の展開深層学習フレームワーク Chainer の開発と今後の展開
深層学習フレームワーク Chainer の開発と今後の展開
 
Chainerの使い方と 自然言語処理への応用
Chainerの使い方と自然言語処理への応用Chainerの使い方と自然言語処理への応用
Chainerの使い方と 自然言語処理への応用
 
東京大学工学部計数工学科応用音響学 D2 Clustering
東京大学工学部計数工学科応用音響学 D2 Clustering東京大学工学部計数工学科応用音響学 D2 Clustering
東京大学工学部計数工学科応用音響学 D2 Clustering
 
2016word embbed
2016word embbed2016word embbed
2016word embbed
 
論文紹介:PaperRobot: Incremental Draft Generation of Scientific Idea
論文紹介:PaperRobot: Incremental Draft Generation of Scientific Idea論文紹介:PaperRobot: Incremental Draft Generation of Scientific Idea
論文紹介:PaperRobot: Incremental Draft Generation of Scientific Idea
 
Jpoug 20120721
Jpoug 20120721Jpoug 20120721
Jpoug 20120721
 
「TensorFlow Tutorialの数学的背景」 クイックツアー(パート1)
「TensorFlow Tutorialの数学的背景」 クイックツアー(パート1)「TensorFlow Tutorialの数学的背景」 クイックツアー(パート1)
「TensorFlow Tutorialの数学的背景」 クイックツアー(パート1)
 
TensorFlow math ja 05 word2vec
TensorFlow math ja 05 word2vecTensorFlow math ja 05 word2vec
TensorFlow math ja 05 word2vec
 
充足可能性問題のいろいろ
充足可能性問題のいろいろ充足可能性問題のいろいろ
充足可能性問題のいろいろ
 
Rubyの御先祖CLUのお話(原本)
Rubyの御先祖CLUのお話(原本)Rubyの御先祖CLUのお話(原本)
Rubyの御先祖CLUのお話(原本)
 
R による文書分類入門
R による文書分類入門R による文書分類入門
R による文書分類入門
 
bigdata2012ml okanohara
bigdata2012ml okanoharabigdata2012ml okanohara
bigdata2012ml okanohara
 
機械学習を用いたWeb上の産学連携関連文書の抽出
機械学習を用いたWeb上の産学連携関連文書の抽出機械学習を用いたWeb上の産学連携関連文書の抽出
機械学習を用いたWeb上の産学連携関連文書の抽出
 
Xtext 紹介
Xtext 紹介Xtext 紹介
Xtext 紹介
 
Hayashi tutorial ne2017
Hayashi tutorial ne2017Hayashi tutorial ne2017
Hayashi tutorial ne2017
 
Tokyo webmining 複雑ネットワークとデータマイニング
Tokyo webmining 複雑ネットワークとデータマイニングTokyo webmining 複雑ネットワークとデータマイニング
Tokyo webmining 複雑ネットワークとデータマイニング
 
Get To The Point: Summarization with Pointer-Generator Networks_acl17_論文紹介
Get To The Point: Summarization with Pointer-Generator Networks_acl17_論文紹介Get To The Point: Summarization with Pointer-Generator Networks_acl17_論文紹介
Get To The Point: Summarization with Pointer-Generator Networks_acl17_論文紹介
 

Proof summit2014mizar

  • 1. Mizarによる形式証明の簡単 な実例 (といろいろ宣伝) Stowarzyszenie Uzytkownikow Mizara (信州大学) ↑ポーランド語でASSOCIATION OF MIZAR USERSらしい @Alcor80UMa (岡崎裕之)
  • 2. Mizar 大熊座(Ura Major略してUma)ζ 星:79番星 なお、Alcorは80番星
  • 3. 目的:Mizarの宣伝 •ProofSummit 2014 定理証明支援系 (Coq,Agda,Isabelle/HOL,ACL2などなど) ↑ここ や自動証明器などを中心に集まるお祭りです。 2014/9/7
  • 4. Formalizing 100 Theorems http://www.cs.ru.nl/~freek/100/ 1.HOL Light 86 2.Mizar 61 3.Isabelle 54 4.Coq 49 5.以下略 全体での達成率89%
  • 5. やっぱりアウェーな感じですが気 にせずやってみましょう お願い: 単なるユーザーですので難しい話は聞かない でください。 (暗号の人なのでこのあたりに居ます。 日本応用数理学会 「数理的技法による情報セキュリティ」研究部会 http://nicosia.is.s.u-tokyo.ac.jp/jsiam-fais/ ) 2014/9/7
  • 6. begin 中正さんのスライド: Mizarプロジェクト概要と現状 日本における参加状況 (スライド提供:信州大学D2中正和久)
  • 7. 形式化数学とは •概念 –理論とは、「記号列を操作するゲーム」である –公理とは、記号列を操作するルールである –定理とは、ルールに則って得られる記号列である –証明とは、定理を得るまでの記号列の操作手順である •歴史 –古典的な形式化数学 •ヒルベルト・プログラム(1920) •ブルバキ「数学原論」(1935-) –計算機支援による形式化数学(1960s-) •Mizar(1973)
  • 8. 近年の背景事情 1.現代数学の高度化・複雑化 –Hilbert~Grothendieckを流れとした数学の抽象化と高度な発展 –数論幾何・超弦理論など高度に発展した分野同士の融合 2.システムの堅牢性への要求の高まり –暗号理論・セキュリティなど、現代数学がシステムの堅牢性を支えている –数学自体の堅牢性を確かめておかなければならない 3.情報科学の成熟と新たな成功 –クラウドコンピューティング(Google) –集合知(Wikipedia) •現代数学の知識体系を計算機で処理可能なコーパスにしようという機運 は確実に高まっている(MKM: Mathematical Knowledge Management) –QED manifesto (1994) –MKM-IG (2001)
  • 9. Mizarプロジェクト 1.Mizar 言語 3.MML 2.Mizar Verifier Mizarで書かれた証明を 検証するプログラム 数学の証明を 記述する言語 証明済みの定義や定理を 集めたライブラリ 特長1. 高い可読性 特長3. 膨大な資産 特長2. 検証に特化 閲覧・引用 検証 投稿
  • 10. Mizarの特長1. 高い可読性 •数学証明記述に特化 •自然言語に近い文法 •一階述語論理 •宣言型言語 •柔軟な型定義 •記号が少ない Freek Wiedijk: Comparing mathematical provers より引用
  • 11. Mizarの特長2. 検証に特化 •他の多くのシステム(HOL, Coq, …)は自動証明に力 点を置いている •Mizarは証明のチェックに力点を置いている –数学知識体系のライブラリ化が主目的 –人間と計算機にそれぞれ得意な分野を分担させる思想。 証明の記述は人間が、検証は計算機が行う。 ませまてぃしゃんの人が趣味で作ったのでこうなった? 集合理論ベース(選択公理)で一階述語論
  • 12. Mizarの特長3. 膨大な資産 •論文誌Formalized Mathematicsは1990年発刊 •11,600の定義と53,700の定理、261万行 (2013年 7月時点) •オープンライセンス(GNU v3/CC BY-SA v3)でWeb 上に公開 –http://mizar.org/library/ •現在も活発に執筆が続けられている
  • 13. Mizarの記述例(TAYLOR展開定理) theorem :: TAYLOR_1:33 for n be Nat, f be PartFunc of REAL,REAL, x0,r be Real st ( 0 < r & f is_differentiable_on n+1, ].x0-r,x0+ r.[ ) for x be Real st x in ].x0-r, x0+r.[ holds ex s be Real st 0 < s & s< 1 & f.x=Partial_Sums(Taylor(f, ].x0-r,x0+r.[,x0,x)).n + (diff(f,].x0- r,x0+r.[).(n+1)).(x0+s*(x-x0)) * (x-x0) |^ (n+1) / ((n+1)!);
  • 14. 他システムで記述すると |- !f diff h n. &0 < h /¥ 0 < n /¥ (diff(0) = f) /¥ (!m t. m < n /¥ &0 <= t /¥ t <= h ==> (diff(m) diffl diff(SUC m)(t))(t)) ==> (?t. &0 < t /¥ t < h /¥ (f(h) = sum(0,n)(¥m. (diff(m)(&0) / &(FACT m)) * (h pow m)) + ((diff(n)(t) / &(FACT n)) * (h pow n)))) 判読が難しい (例 HOL Light)
  • 15. 膨大な数学定理の形式証明ライブラリ を作るには •どの定理や定義をどのように並べ,有機的に関 連づけていくのか? •どのように形式証明を記述していくか? •組織立てられた作成計画が必須 •デザインなど人間の高次な脳の機能に依存 •視野が観ている窓に限定される対話型の定理 証明支援系では困難
  • 16. 現代数学の大規模なライブラリ構築 に重点をおくなら 現状の計算機の技術では, •形式証明自体は人間が書く •計算機はそれを精査する –それぞれ得意なものに集中する役割分業が有効 •QEDの目的達成にはほとんどの数理・情報科学の 研究者が,自分が執筆する通常の論文と同様に, 読み書きに不自由しないシステムの開発が重要
  • 17. MML(mizar mathematical library) music macro languageではない 2013年6月時点で, Mizar Mathematical Library(MML)は •延べ200名を超える著者 •1181のファイル •約52,000の定理 •10,000を超える形式上の定義 WEB上に公開 http://mizar.uwb.edu.pl/version/current/html/
  • 18. BiaLystok大学(ポーランド), Alberta大学(カナダ), 信州大学(日本)で研究グループによって展開, 推進されている http://mizar.org/project/ チュートリアル, ライブラリ検索, 各種リソース markun.cs.shinshu-u.ac.jp/kiso/projects/proofchecker/mizar/ index-j.html 成果と資源はweb上ですべて公開
  • 19. 形式化記述の例(環境設定部分) environ Vocabularies NUMBERS, INT_1, ORDINAL1, SUBSET_1, CARD_1, ARYTM_3, ARYTM_1, RELAT_1, NAT_1,XCMPLX_0,REAL_1; Notations SUBSET_1,ORDINAL1,NUMBERS, XCMPLX_0, INT_1, NAT_1; constructors REAL_1,NAT_1, INT_1, CARD_1; Registrations XREAL_0,NAT_1, INT_1, ORDINAL1, XBOOLE_0; Requirements REAL, NUMERALS, SUBSET,ARITHM, BOOLE; definitions INT_1; Theorems INT_1; Schemes NAT_1;
  • 20. 環境部には •Vocabularies 用語の定義ファイルを指定する。 •Notations •constructors 個々の概念の定義する 概念間の階層構造等の関係を表す。 •registrations 変数型の階層的な関係,数学的属性の従属関係などを指定する。 •Requirements 実数の演算や集合の包含関係の演算の自動的な処理を指定する。 •definitions •theorems •schemes begin文以下の形式化記述の本体で引用される既存の定義や定理, 公理図式の収録ファイルを指定する。
  • 21. 形式化記述の例(本体部分) begin theorem EX1: for n being Nat holds 2 divides n*(n-1) proof defpred P[Nat] means 2 divides $1*($1-1); 0 = 2 * 0; then P0: P[0] by INT_1:def 3; PN: for n being Nat st P[n] holds P[n+1] proof let n be Nat; assume P[n]; then consider s being Integer such that A2: n * (n - 1) = 2 * s by INT_1:def 3 (n+1)*(n+1-1) = n*(n-1) + n + n*1 .=n*(n-1) + 2*n .=2*s + 2*n by A2 .= 2*(s+n); hence P[n+1] by INT_1:def 3; end; thus for n being Nat holds P[n] from NAT_1:sch 2(P0,PN); end;
  • 22. theorem EX1: 命題名EX1という命題のラベル名の宣 言 for n being Nat holds 2 divides n*(n-1) 「任意の自然数nに対して,n・(n-1)は 2で割り切れる。」という命題 (注) Natは自然数の変数型名であ る。 proof 命題EX1の証明開始 (注) 以下proofとend;で挟まれた proof ~ endブロックが書かれ,上記 命題EX1の証明が記述される。 defpred P[Nat] means 2 divides $1*($1-1); 自然数を引数に持つ述語P [ ] で,引 数($1)と$1-1との積($1・($1-1))が,2で 割り切れることを表わす述語を定義 する。(defpred) 0 = 2 * 0; then P0: P[0] by INT_1:def 3; 0 = 2・0; だから 2 divides 0・(0-1); であり 補題名P0 : P[0] が 定義INT_1:def 3 により成り立っている。
  • 23. PN: for n being Nat st P[n] holds P[n+1] 補題PN: 任意の自然数nに対して「P[n]が成り 立つときP[n+1]が成り立つ」 proof 補題PNの証明開始 proof ~ endブロックで記述する。 let n be Nat; nを任意の自然数とする assume P[n]; P[n] すなわち2 divides n・(n-1) を仮定する then consider s being Integer such that A2: n * (n – 1) = 2 * s by INT_1:def 3; 従って INT_1:def 3によれば以下の補 題A2が成り立つある自然数 s を考え ることができる。 補題A2: n・(n - 1) = 2・s
  • 24. (n+1)*(n+1-1) = n*(n-1) + n + n*1 .=n*(n-1) + 2*n .=2*s + 2*n by A2 .= 2*(s+n); 補題2などにより,等式 (n+1)・(n+1-1) = 2・(s+n) が得られる hence P[n+1] by INT_1:def 3; よってP[n+1] がINT_1:def 3によって成 り立つ end; 補題PNの証明終了 thus for n being Nat holds P[n] from NAT_1:sch 2(P0,PN); 以上補題P0,PNと数学的帰納法の公理 図式NAT_1:sch 2を用いて, 任意の 自然数nについてP[n]が成り立つ end; 定理EX1の証明終了
  • 25. 英文に近い表現の記述を眺めれば Natが自然数を表すものと分かれば 『任意の自然数nについて,2はn・(n-1)を割り 切る』という命題とそれの数学的帰納法による 証明を記述していることが大凡判読できる Mizarは自然言語に近い表現ができることが特徴
  • 26. Mizarライブラリ(MML)の現状 •2013/06/18現在で1181の,数学ライブラリが整備されている •システムとライブラリは,インターネット上で全て公開されている。 http://Mizar.uwb.edu.pl/version/current/html/ •数学の公理体系から出発して必要な定理に至る全ての結果・証明が完 全に収録され,しかも他を参照する必要のない自己完結的な形式化数 学記述のライブラリとして提供されている。 •WEB表示中のproof という単語をクリックすれば,証明の全文が表示さ れる。さらにそれで引用されるライブラリ中の既存の定理や用語にもリン クが張られ,これらのリンクを繰り返し辿れば,集合論の公理系にまで戻 ることができる。
  • 27. 情報工学関連(日本側が始めた活動) DIST_* 離散集合上の確率変数 DESCIP_1,AESCIP_1 秘密鍵暗号アルゴリズム CIRCUIT*,ftacell1, gfacirc* 回路,並列回路 PERTRI* ペトリネット morph_01 画像処理と数理形態学 NTALOG_1 拡張ユーグリッド互除法 ZMOD_* ECPF_* 整数環上の Module,楕円 曲線 AMI_*,SCM* 抽象計算機
  • 28. WEB上で利用できるMizarシステム(1) •主に修士の学生の講義用に使用している。 •新しく定義するか特定のシステムコマンドを使用するかどうかによって Verifierのパッケージを選択 •次の入力画面で使用する学生のIDやe-mailアドレスを入力する。 •Verifierのメインページが表示される。 証明を記入するテキストボックス ファイル名などを記入する入力ボックス がある。 • データを入力した後SUBMITボタンをマウスでクリックすると,Mizarの proof checker が働いて, 証明をチェックし誤りがあれば,エラーメッセー ジが表示される •とはいってもインストールも簡単。Install.bat実行して 環境変数2つ(MIZFILES,MIZPATH)追加するだけ。 http://www.wakasato.jp/Mizar/
  • 29. WEB上で利用できるMizarシステム(2) •Mizarの開発者の一人であるGrzegorz Bancerek 教授がJSPSの客員研究として本学 に招聘した際に作成 •ライブラリの高度な検索機能 http://mmlquery.Mizar.org/mmlquery/three.html
  • 30. WEB上で利用できるMizarシステム(3) •Mizar Projectの重要なメンバーであるJosef Urban教授が開 発 •証明記述中の既存のMMLの用語,引用定理へ全てリンクが 張られている •E-learning教材を開発する際には,他を参照する必要のない 自己完結的な形式化数学記述のライブラリとして提供できる •Mizarの既存のライブラリと一体で数学の公理体系から出発 して必要な定理に至る全ての結果・証明を完全に収録した教 材が作成できる http://Mizar.cs.ualberta.ca/~mptp/Mizar.html
  • 31. 終わりに Mizarの概要 プロジェクトの現状と日本おける参加状況 数理的思考訓練用の教材を汎用 CMSに組み込む研究 e-learningを主体とした社会人向け のインターネットによる遠隔教育 暗号理論や情報工学分野での 形式検証に必要な ライブラリの整備を当面の目標に ライブラリ作成に取り組んでいる. 学部や修士学生が学習する微積分や関数解析, 代数等の基礎的定理のライブラリ整備
  • 32. •e-learning教材を開発する際に他を 参照する必要のない自己完結的な 教材を提供できる. •Mizarの既存のライブラリと一体で 数学の公理体系から出発して必要 な定理に至る全ての結果・ 証明 を収録した教材の作成ができる. 数理的思考訓練用の 教材を汎用CMSに組 み込む研究 e-learningを主体とし た社会人向けのイン ターネットによる遠隔 教育 QED manifestoの目的の一つである 数学的な思考能力の開発に合致する. 終わりに
  • 33. 見果てぬ夢 •殆どの数学定理のライブラリ化が完成し,数学,数理科学, 応用数学分野の研究者が,形式化数学記述システムを用い て自分たちの最新の研究成果とその証明を記述できるよう になることを期待している. •証明の正しさは,どのような無名な研究者の論文であっても, 学会の権威者ではなく,計算機が瞬時に判定する. •従来の人間の査読は人間系によってのみ可能な結果の有 用性,新規性といったより高度な判定に利用される. •プロ,アマチュアの境がなくなり,学問研究の数理科学分野 でのユビキタスが実現されることを期待している. end 中正さんのスライド;
  • 34. 簡単な例をもう一つ theorem for x be Element of NAT st 1 < x holds not ex N,c be Element of NAT st for n be Element of NAT st N <= n holds x to_power n <= c * ( n to_power x) 2014/9/7 xnncxholdsnNtsntsNcholdsxtsx     ...., 1.. NNN
  • 35. つまりこれを言いたい 2014/9/7 ( ) n x x Ο n 計算量の評価 とか 暗号の安全性の議論に使う
  • 36. theorem N2POWINPOLY: for x be Element of NAT st 1 < x holds not ex N,c be Element of NAT st for n be Element of NAT st N <= n holds 2 to_power n <= c * ( n to_power x) を使って for x be Element of NAT st 1 < x holds not ex N,c be Element of NAT st for n be Element of NAT st N <= n holds x to_power n <= c * ( n to_power x) を証明 2014/9/7
  • 37. proof let x be Element of NAT; assume AS: 1 < x; assume CNT: ex N,c be Element of NAT st for n be Element of NAT st N <= n holds x to_power n <= c * ( n to_power x); まず命題が偽であると仮定する 2014/9/7
  • 38. ex N,c be Element of NAT st for n be Element of NAT st N <= n holds 2 to_power n <= c * ( n to_power x); (proofは次ページに) hence contradiction by AS,N2POWINPOLY; end; 変なことが証明できたので定理N2POWINPOLY に矛盾して背理法で証明終わり 2014/9/7
  • 39. ex N,c be Element of NAT st for n be Element of NAT st N <= n holds 2 to_power n <= c * ( n to_power x) proof consider N,c be Element of NAT such that CNT2: for n be Element of NAT st N <= n holds x to_power n <= c * ( n to_power x) by CNT; take N,c; for n be Element of NAT st N <= n holds 2 to_power n <= c * ( n to_power x) proof let n be Element of NAT; assume N <= n;then LCX1: x to_power n <= c * ( n to_power x) by CNT2; 1+1 <= x by AS,INT_1:7;then 2 to_power n <= x to_power n by LEMC01; hence thesis by LCX1,XXREAL_0:2; end; hence thesis; end; 2014/9/7
  • 40. 数学の証明にしか使えない? •そういうわけでもない –アルゴリズム –回路の検証 等にも使えます。 例えば暗号とか 2014/9/7
  • 41. Data Encryption Standard (DES) Data Encryption Standard (DES) Symmetric cryptosystem (Block cipher). Selected by the National Bureau of Standards as an official Federal Information Processing Standard for the United States in 1976 (FIPS46). 3 Strong influence on the design of its successors. Insecure Advanced Encryption Standard (AES) Now
  • 42. Formal Verification of DES Using the Mizar Proof Checker Hiroyuki Okazaki1, Kenichi Arai2, Yasunari Shidama1 1. Shinshu University 2. Nagano Technical High School(当時) 現:東京理科大学 July 18, 2011
  • 43. Agenda Proof is verified by using a proof checker. The correctness of our formalization of the DES algorithm. Prove the security of cryptographic systems by using the Mizar proof checker. Introduction : Our formalization of the Data Encryption Standard (DES) algorithm. 1
  • 44. Proof Checker “Mizar” The definitions and the theorems have been verified for correctness using the Mizar proof checking system. Mizar Mizar is an advanced project by Mizar Society that Prof. A. Trybulec leads in Bialystok university to which is formalized mathematics by the computer-aid. The Mizar project describes mathematical proofs in the Mizar language, which is created to formally describe mathematics. What formalizes the proof of mathematics and describes it is called article. 2
  • 45. 4 Structure of DES Figure 1: Structure of DES 64bits length plaintext block. 64bits length secret key. 64bits length ciphertext block. 16 rounds of processing iterations Decryption Encryption Same Key Feistel Structure 64 bits 64 bits 64 bits
  • 46. i-th Round of Feistel Structure 5 Figure 2: i-th round of Feistel structure ( ) Ri-1 : 32 bits length block. Li-1 : 32 bits length block. Ki : i-th round key (48 bits). Exclusive OR (XOR)
  • 47. About “functor” and “Function” 6 Mizar Two ways to define computational routines in an algorithmic sense. functor function A “functor” is a relation between the input and output of a routine. A “function” is a map from the space of the input onto that of the output.
  • 48. Strategy of Formalizing DES in Mizar 7 Step1 : Formalization of the algorithm of generalized DES. Step2 : Formalization of the primitives of DES according to FIPS46–3. Formalization of the DES algorithm
  • 49. Formalization of the generalized algorithm of DES 8 Definition 5.1: ( Codec of gereralized DES ) Mizar language it = DES-like- CoDec(M,F,IP,RK) : Concatenation : Inverse [:A, B:] : Cartesian product of A and B
  • 50. Correctness of the generalized algorithm of DES Theorem 5.1: ( Correctness of generalized DES ) 9 Proved that the ciphertext encoded by any Feistel cipher algorithm can be decoded uniquely with the same algorithm and secret key that were used in encryption. Mizar language
  • 51. Formalization of DES Formalization of the DES algorithm according to FIPS46–3 in Mizar language. Step1: Formalization of the DES primitives according to FIPS46-3. Step2: Formalization of the correctness of the DES algorithm. (Using the formalization of the generalized DES algorithm.) Prove the correctness of the DES algorithm 10
  • 52. 11 Formalization of DES Primitives S-Boxes ( S-Box S1, ... , S-Box S5, .. , S-Box S8 ). Initial Permutation (IP). Final Permutation (IP -1). Feistel Function : ( ). Bit selection Function (E). Permutation (P). Key Scheduling Function (KS). Permuted Choice 1 (PC1). Permuted Choice 2 (PC2). Left Shift. DES Primitives
  • 53. S-Boxes 12 S1 S2 … S7 S8 …
  • 54. 13 S-Boxes Theorem 6.1: ( S-Box S1) Similarly defined the other S-Boxes, DES-SBOX2 (S2), ….. , DES-SBOX8 (S8). S-Box S1 Mizar language it = DES-SBOX1
  • 55. 14 Initial Permutation (IP) Definition 6.2: ( IP as functor ) IP Mizar language it = DES-IP(r)
  • 56. Initial Permutation (IP) 15 Definition 6.3: ( IP as function ) Similarly defined the functor of the final permutation DES-IPINV and the function of the DES-PIPINV. Mizar language it = DES-PIP
  • 57. 16 Feistel function Figure 4: Feistel function Feistel Function ( ). Bit selection Function (E) Permutation (P) S-Boxes + + 32 bits 48 bits 48 bits 48 bits 6 bits 6 bits 6 bits 6 bits 6 bits 6 bits 6 bits 6 bits 4 bits 4 bits 4 bits 4 bits 4 bits 4 bits 4 bits 4 bits 32 bits 32 bits
  • 58. 17 Bit Selection Function (E) E Definition 6.4: ( E as functor ) Mizar language it = DES-E(r)
  • 59. Permutation (P) 18 P Definition 6.5: ( P as functor ) Mizar language it = DES-P(r)
  • 60. Formalization of Feistel Function Definition 6.6: ( Feistel function as functor ) 19 Mizar language it = DES-F(R, Rkey)
  • 61. 20 Formalization of Feistel Function Definition 6.7: ( Feistel function as function ) Mizar language it = DES-FFUNC
  • 62. 21 Key Scheduling Function (KS) Figure 5: Key Scheduling Function Key Scheduling Function (KS). Permuted Choice 1 (PC1) Permuted Choice 2 (PC2) Left Shift + + Kn = KS(n, Secret key) 64 bits 28 bits 28 bits 48 bits 56 bits 28 bits 28 bits 28 bits 28 bits 28 bits 28 bits 48 bits 56 bits
  • 63. Permuted Choice 1 (PC1) Definition 6.8: ( PC1 as functor ) Mizar language it = DES-PC1(r) PC1 Similarly defined the functor of PC2 as DES-PC2. 22 C0 D0
  • 64. 23 Table of the Numbers of Left-Shift Definition 6.9: ( Table of Left-Shift ) Mizar language it = bitshift_DES i-th Numbers Round of Left-Shift 2 1 3 2 4 2 5 2 6 2 7 2 8 2 9 1 10 2 11 2 12 2 13 2 14 2 15 2 16 1 1 1
  • 65. 24 Formalization of the Key Scheduling Function Definition 6.7: ( Key Scheduling function ) Mizar language it = DES-KS(Key)
  • 66. 25 DES Algorithm The generalized DES algorithm The DES primitives Definition 6.11: ( DES Algorithm ) Mizar language it = DES-CoDEC (M,F,IP,RK) DES algorithm
  • 67. 26 Encode and Decode Algorithm of DES Definition 6.12: ( Encode Algorithm of DES) Mizar language Definition 6.13: ( Decode Algorithm of DES) Mizar language
  • 68. Correctness of DES Theorem 6.1: ( Correctness of DES ) Proved using the Mizar system that the ciphertext encoded by the DES algorithm can be decoded uniquely with the same algorithm and secret key that were used in Encryption. 27 Mizar language
  • 69. Conclusion 28 Prove the correctness of the DES algorithm. Introduction : Our formalization of DES algorithm. Mizar proof checking system. Future work Analyze the security of DES. Prove the security of cryptographic systems by using the Mizar proof checker.
  • 70. TPP告知 TPP2014もよろしくね! (Theorem Proving and Provers Meeting) 日時:2014 年 12 月 3 日(水) 〜 12 月 5 日(金) 場所:九州大学・西新プラザ (福岡市早良区西新 2-16-23) 講演者募集中! Adam Chlipala 氏 (MIT),Cyril Cohen氏 (Univ.Gothenburg) をお 招きしてCoq および Ssreflect,MathComp に関する御講演を 頂く予定だそうです。 入口にビラが置いてありますのでよろしければお持ちください http://coop-math.ism.ac.jp/files/143/Announcement2.pdf 2014/9/7