| author | wenzelm | 
| Fri, 20 Sep 2024 23:37:00 +0200 | |
| changeset 80917 | 2a77bc3b4eac | 
| parent 80914 | d97fdabd9e2b | 
| permissions | -rw-r--r-- | 
| 
40967
 
5eb59b62e7de
avoid explicit encoding -- acknowledge UTF-8 as global default and Isabelle/jEdit preference of UTF-8-Isabelle;
 
wenzelm 
parents: 
39246 
diff
changeset
 | 
1  | 
(* Author: Ning Zhang and Christian Urban  | 
| 17505 | 2  | 
|
3  | 
Example theory involving Unicode characters (UTF-8 encoding) -- both  | 
|
4  | 
formal and informal ones.  | 
|
5  | 
*)  | 
|
6  | 
||
| 61343 | 7  | 
section \<open>A Chinese theory\<close>  | 
| 17505 | 8  | 
|
| 17553 | 9  | 
theory Chinese imports Main begin  | 
| 17505 | 10  | 
|
| 61343 | 11  | 
text\<open>数学家能把咖啡变成理论,如今中国的数学家也可  | 
| 17505 | 12  | 
以在伊莎贝拉的帮助下把茶变成理论.  | 
13  | 
||
14  | 
伊莎贝拉-世界数学家的新宠,现今能识别中文,成为  | 
|
15  | 
中国数学家理论推导的得力助手.  | 
|
16  | 
||
| 61343 | 17  | 
\<close>  | 
| 17505 | 18  | 
|
| 58310 | 19  | 
datatype shuzi =  | 
| 
80914
 
d97fdabd9e2b
standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
 
wenzelm 
parents: 
61343 
diff
changeset
 | 
20  | 
One (\<open>一\<close>)  | 
| 
 
d97fdabd9e2b
standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
 
wenzelm 
parents: 
61343 
diff
changeset
 | 
21  | 
| Two (\<open>二\<close>)  | 
| 
 
d97fdabd9e2b
standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
 
wenzelm 
parents: 
61343 
diff
changeset
 | 
22  | 
| Three (\<open>三\<close>)  | 
| 
 
d97fdabd9e2b
standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
 
wenzelm 
parents: 
61343 
diff
changeset
 | 
23  | 
| Four (\<open>四\<close>)  | 
| 
 
d97fdabd9e2b
standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
 
wenzelm 
parents: 
61343 
diff
changeset
 | 
24  | 
| Five (\<open>五\<close>)  | 
| 
 
d97fdabd9e2b
standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
 
wenzelm 
parents: 
61343 
diff
changeset
 | 
25  | 
| Six (\<open>六\<close>)  | 
| 
 
d97fdabd9e2b
standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
 
wenzelm 
parents: 
61343 
diff
changeset
 | 
26  | 
| Seven (\<open>七\<close>)  | 
| 
 
d97fdabd9e2b
standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
 
wenzelm 
parents: 
61343 
diff
changeset
 | 
27  | 
| Eight (\<open>八\<close>)  | 
| 
 
d97fdabd9e2b
standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
 
wenzelm 
parents: 
61343 
diff
changeset
 | 
28  | 
| Nine (\<open>九\<close>)  | 
| 
 
d97fdabd9e2b
standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
 
wenzelm 
parents: 
61343 
diff
changeset
 | 
29  | 
| Ten (\<open>十\<close>)  | 
| 17505 | 30  | 
|
| 
80914
 
d97fdabd9e2b
standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
 
wenzelm 
parents: 
61343 
diff
changeset
 | 
31  | 
primrec translate :: "shuzi \<Rightarrow> nat" (\<open>|_|\<close> [100] 100) where  | 
| 17505 | 32  | 
"|一| = 1"  | 
| 39246 | 33  | 
|"|二| = 2"  | 
34  | 
|"|三| = 3"  | 
|
35  | 
|"|四| = 4"  | 
|
36  | 
|"|五| = 5"  | 
|
37  | 
|"|六| = 6"  | 
|
38  | 
|"|七| = 7"  | 
|
39  | 
|"|八| = 8"  | 
|
40  | 
|"|九| = 9"  | 
|
41  | 
|"|十| = 10"  | 
|
| 17505 | 42  | 
|
43  | 
thm translate.simps  | 
|
44  | 
||
45  | 
lemma "|四| + |六| = |十|"  | 
|
46  | 
by simp  | 
|
47  | 
||
48  | 
end  |