| author | wenzelm | 
| Tue, 08 Oct 2024 12:10:35 +0200 | |
| changeset 81125 | ec121999a9cb | 
| 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: 
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 = | 
| 80914 
d97fdabd9e2b
standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
 wenzelm parents: 
62812diff
changeset | 20 | Alef (\<open>א\<close>) | 
| 
d97fdabd9e2b
standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
 wenzelm parents: 
62812diff
changeset | 21 | | Bet (\<open>ב\<close>) | 
| 
d97fdabd9e2b
standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
 wenzelm parents: 
62812diff
changeset | 22 | | Gimel (\<open>ג\<close>) | 
| 
d97fdabd9e2b
standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
 wenzelm parents: 
62812diff
changeset | 23 | | Dalet (\<open>ד\<close>) | 
| 
d97fdabd9e2b
standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
 wenzelm parents: 
62812diff
changeset | 24 | | He (\<open>ה\<close>) | 
| 
d97fdabd9e2b
standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
 wenzelm parents: 
62812diff
changeset | 25 | | Vav (\<open>ו\<close>) | 
| 
d97fdabd9e2b
standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
 wenzelm parents: 
62812diff
changeset | 26 | | Zayin (\<open>ז\<close>) | 
| 
d97fdabd9e2b
standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
 wenzelm parents: 
62812diff
changeset | 27 | | Het (\<open>ח\<close>) | 
| 
d97fdabd9e2b
standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
 wenzelm parents: 
62812diff
changeset | 28 | | Tet (\<open>ט\<close>) | 
| 
d97fdabd9e2b
standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
 wenzelm parents: 
62812diff
changeset | 29 | | Yod (\<open>י\<close>) | 
| 
d97fdabd9e2b
standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
 wenzelm parents: 
62812diff
changeset | 30 | | Kaf (\<open>כ\<close>) | 
| 
d97fdabd9e2b
standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
 wenzelm parents: 
62812diff
changeset | 31 | | Lamed (\<open>ל\<close>) | 
| 
d97fdabd9e2b
standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
 wenzelm parents: 
62812diff
changeset | 32 | | Mem (\<open>מ\<close>) | 
| 
d97fdabd9e2b
standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
 wenzelm parents: 
62812diff
changeset | 33 | | Nun (\<open>נ\<close>) | 
| 
d97fdabd9e2b
standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
 wenzelm parents: 
62812diff
changeset | 34 | | Samekh (\<open>ס\<close>) | 
| 
d97fdabd9e2b
standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
 wenzelm parents: 
62812diff
changeset | 35 | | Ayin (\<open>ע\<close>) | 
| 
d97fdabd9e2b
standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
 wenzelm parents: 
62812diff
changeset | 36 | | Pe (\<open>פ\<close>) | 
| 
d97fdabd9e2b
standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
 wenzelm parents: 
62812diff
changeset | 37 | | Tsadi (\<open>צ\<close>) | 
| 
d97fdabd9e2b
standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
 wenzelm parents: 
62812diff
changeset | 38 | | Qof (\<open>ק\<close>) | 
| 
d97fdabd9e2b
standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
 wenzelm parents: 
62812diff
changeset | 39 | | Resh (\<open>ר\<close>) | 
| 
d97fdabd9e2b
standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
 wenzelm parents: 
62812diff
changeset | 40 | | Shin (\<open>ש\<close>) | 
| 
d97fdabd9e2b
standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
 wenzelm parents: 
62812diff
changeset | 41 | | Tav (\<open>ת\<close>) | 
| 17414 | 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 |