| author | Lars Hupel <lars.hupel@mytum.de> | 
| Tue, 05 Jul 2016 09:50:20 +0200 | |
| changeset 63386 | 0d6adf2d5035 | 
| parent 62812 | ce22e5c3d4ce | 
| child 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: 
39246diff
changeset | 1 | (* Author: Makarius | 
| 17414 | 2 | |
| 3 | Example theory involving Unicode characters (UTF-8 encoding) -- both | |
| 4 | formal and informal ones. | |
| 5 | *) | |
| 6 | ||
| 59031 | 7 | section \<open>A Hebrew theory\<close> | 
| 17506 | 8 | |
| 17414 | 9 | theory Hebrew | 
| 10 | imports Main | |
| 11 | begin | |
| 12 | ||
| 62809 
4b8f08de2792
explicit warning about bidi uncertainty in Unicode;
 wenzelm parents: 
59031diff
changeset | 13 | text \<open> | 
| 62812 
ce22e5c3d4ce
more robust display of bidirectional Unicode text: enforce left-to-right;
 wenzelm parents: 
62809diff
changeset | 14 | \<^bold>\<open>Warning:\<close> Bidirectional Unicode text may confuse display in browsers, editors, etc.! | 
| 62809 
4b8f08de2792
explicit warning about bidi uncertainty in Unicode;
 wenzelm parents: 
59031diff
changeset | 15 | \<close> | 
| 
4b8f08de2792
explicit warning about bidi uncertainty in Unicode;
 wenzelm parents: 
59031diff
changeset | 16 | |
| 
4b8f08de2792
explicit warning about bidi uncertainty in Unicode;
 wenzelm parents: 
59031diff
changeset | 17 | subsection \<open>The Hebrew Alef-Bet (א-ב).\<close> | 
| 17414 | 18 | |
| 58310 | 19 | datatype alef_bet = | 
| 17414 | 20 |     Alef    ("א")
 | 
| 21 |   | Bet     ("ב")
 | |
| 22 |   | Gimel   ("ג")
 | |
| 23 |   | Dalet   ("ד")
 | |
| 24 |   | He      ("ה")
 | |
| 25 |   | Vav     ("ו")
 | |
| 26 |   | Zayin   ("ז")
 | |
| 27 |   | Het     ("ח")
 | |
| 28 |   | Tet     ("ט")
 | |
| 29 |   | Yod     ("י")
 | |
| 30 |   | Kaf     ("כ")
 | |
| 31 |   | Lamed   ("ל")
 | |
| 32 |   | Mem     ("מ")
 | |
| 33 |   | Nun     ("נ")
 | |
| 34 |   | Samekh  ("ס")
 | |
| 35 |   | Ayin    ("ע")
 | |
| 36 |   | Pe      ("פ")
 | |
| 37 |   | Tsadi   ("צ")
 | |
| 38 |   | Qof     ("ק")
 | |
| 39 |   | Resh    ("ר")
 | |
| 40 |   | Shin    ("ש")
 | |
| 41 |   | Tav     ("ת")
 | |
| 42 | ||
| 43 | thm alef_bet.induct | |
| 44 | ||
| 45 | ||
| 62809 
4b8f08de2792
explicit warning about bidi uncertainty in Unicode;
 wenzelm parents: 
59031diff
changeset | 46 | subsection \<open>Interpreting Hebrew letters as numbers.\<close> | 
| 17414 | 47 | |
| 62809 
4b8f08de2792
explicit warning about bidi uncertainty in Unicode;
 wenzelm parents: 
59031diff
changeset | 48 | primrec mispar :: "alef_bet \<Rightarrow> nat" | 
| 59031 | 49 | where | 
| 17414 | 50 | "mispar א = 1" | 
| 39246 | 51 | | "mispar ב = 2" | 
| 52 | | "mispar ג = 3" | |
| 53 | | "mispar ד = 4" | |
| 54 | | "mispar ה = 5" | |
| 55 | | "mispar ו = 6" | |
| 56 | | "mispar ז = 7" | |
| 57 | | "mispar ח = 8" | |
| 58 | | "mispar ט = 9" | |
| 59 | | "mispar י = 10" | |
| 60 | | "mispar כ = 20" | |
| 61 | | "mispar ל = 30" | |
| 62 | | "mispar מ = 40" | |
| 63 | | "mispar נ = 50" | |
| 64 | | "mispar ס = 60" | |
| 65 | | "mispar ע = 70" | |
| 66 | | "mispar פ = 80" | |
| 67 | | "mispar צ = 90" | |
| 68 | | "mispar ק = 100" | |
| 69 | | "mispar ר = 200" | |
| 70 | | "mispar ש = 300" | |
| 71 | | "mispar ת = 400" | |
| 17414 | 72 | |
| 73 | thm mispar.simps | |
| 74 | ||
| 75 | lemma "mispar ק + mispar ל + mispar ה = 135" | |
| 76 | by simp | |
| 77 | ||
| 78 | end |