| 81811 |      1 | (*  Title:      HOL/Library/Suc_Notation.thy
 | 
|  |      2 |     Author:     Manuel Eberl and Tobias Nipkow
 | 
|  |      3 | 
 | 
|  |      4 | Compact notation for nested \<open>Suc\<close> terms. Just notation. Use standard numerals for large numbers.
 | 
|  |      5 | *)
 | 
|  |      6 | 
 | 
|  |      7 | theory Suc_Notation
 | 
|  |      8 | imports Main
 | 
|  |      9 | begin
 | 
|  |     10 | 
 | 
|  |     11 | text \<open>Nested \<open>Suc\<close> terms of depth \<open>2 \<le> n \<le> 9\<close> are abbreviated with new notations \<open>Suc\<^sup>n\<close>:\<close>
 | 
|  |     12 | 
 | 
|  |     13 | abbreviation (input) Suc2 where "Suc2 n \<equiv> Suc (Suc n)"
 | 
|  |     14 | abbreviation (input) Suc3 where "Suc3 n \<equiv> Suc (Suc2 n)"
 | 
|  |     15 | abbreviation (input) Suc4 where "Suc4 n \<equiv> Suc (Suc3 n)"
 | 
|  |     16 | abbreviation (input) Suc5 where "Suc5 n \<equiv> Suc (Suc4 n)"
 | 
|  |     17 | abbreviation (input) Suc6 where "Suc6 n \<equiv> Suc (Suc5 n)"
 | 
|  |     18 | abbreviation (input) Suc7 where "Suc7 n \<equiv> Suc (Suc6 n)"
 | 
|  |     19 | abbreviation (input) Suc8 where "Suc8 n \<equiv> Suc (Suc7 n)"
 | 
|  |     20 | abbreviation (input) Suc9 where "Suc9 n \<equiv> Suc (Suc8 n)"
 | 
|  |     21 | 
 | 
|  |     22 | notation Suc2 ("Suc\<^sup>2")
 | 
|  |     23 | notation Suc3 ("Suc\<^sup>3")
 | 
|  |     24 | notation Suc4 ("Suc\<^sup>4")
 | 
|  |     25 | notation Suc5 ("Suc\<^sup>5")
 | 
|  |     26 | notation Suc6 ("Suc\<^sup>6")
 | 
|  |     27 | notation Suc7 ("Suc\<^sup>7")
 | 
|  |     28 | notation Suc8 ("Suc\<^sup>8")
 | 
|  |     29 | notation Suc9 ("Suc\<^sup>9")
 | 
|  |     30 | 
 | 
|  |     31 | text \<open>Beyond 9, the syntax \<open>Suc\<^bsup>n\<^esup>\<close> kicks in:\<close>
 | 
|  |     32 | 
 | 
|  |     33 | syntax
 | 
|  |     34 |   "_Suc_tower" :: "num_token \<Rightarrow> nat \<Rightarrow> nat"  ("Suc\<^bsup>_\<^esup>")
 | 
|  |     35 | 
 | 
|  |     36 | parse_translation \<open>
 | 
|  |     37 |   let
 | 
|  |     38 |     fun mk_sucs_aux 0 t = t
 | 
|  |     39 |       | mk_sucs_aux n t = mk_sucs_aux (n - 1) (\<^const>\<open>Suc\<close> $ t)
 | 
|  |     40 |     fun mk_sucs n = Abs("n", \<^typ>\<open>nat\<close>, mk_sucs_aux n (Bound 0))
 | 
|  |     41 | 
 | 
|  |     42 |     fun Suc_tr [Free (str, _)] = mk_sucs (the (Int.fromString str))
 | 
|  |     43 | 
 | 
|  |     44 |   in [(\<^syntax_const>\<open>_Suc_tower\<close>, K Suc_tr)] end
 | 
|  |     45 | \<close>
 | 
|  |     46 | 
 | 
|  |     47 | print_translation \<open>
 | 
|  |     48 |   let
 | 
|  |     49 |     val digit_consts =
 | 
|  |     50 |         [\<^const_syntax>\<open>Suc2\<close>, \<^const_syntax>\<open>Suc3\<close>, \<^const_syntax>\<open>Suc4\<close>, \<^const_syntax>\<open>Suc5\<close>,
 | 
|  |     51 |          \<^const_syntax>\<open>Suc6\<close>, \<^const_syntax>\<open>Suc7\<close>, \<^const_syntax>\<open>Suc8\<close>, \<^const_syntax>\<open>Suc9\<close>]
 | 
|  |     52 |     val num_token_T = Simple_Syntax.read_typ "num_token"
 | 
|  |     53 |     val T = num_token_T --> HOLogic.natT --> HOLogic.natT
 | 
|  |     54 |     fun mk_num_token n = Free (Int.toString n, num_token_T)
 | 
|  |     55 |     fun dest_Suc_tower (Const (\<^const_syntax>\<open>Suc\<close>, _) $ t) acc =
 | 
|  |     56 |           dest_Suc_tower t (acc + 1)
 | 
|  |     57 |       | dest_Suc_tower t acc = (t, acc)
 | 
|  |     58 | 
 | 
|  |     59 |     fun Suc_tr' [t] =
 | 
|  |     60 |       let
 | 
|  |     61 |         val (t', n) = dest_Suc_tower t 1
 | 
|  |     62 |       in
 | 
|  |     63 |         if n > 9 then
 | 
|  |     64 |           Const (\<^syntax_const>\<open>_Suc_tower\<close>, T) $ mk_num_token n $ t'
 | 
|  |     65 |         else if n > 1 then
 | 
|  |     66 |           Const (List.nth (digit_consts, n - 2), T) $ t'
 | 
|  |     67 |         else
 | 
|  |     68 |           raise Match
 | 
|  |     69 |       end
 | 
|  |     70 | 
 | 
|  |     71 |   in [(\<^const_syntax>\<open>Suc\<close>, K Suc_tr')]
 | 
|  |     72 |  end
 | 
|  |     73 | \<close>
 | 
|  |     74 | 
 | 
|  |     75 | (* Tests
 | 
|  |     76 | 
 | 
|  |     77 | ML \<open>
 | 
|  |     78 | local
 | 
|  |     79 |   fun mk 0 = \<^term>\<open>x :: nat\<close>
 | 
|  |     80 |     | mk n = \<^const>\<open>Suc\<close> $ mk (n - 1)
 | 
|  |     81 |   val ctxt' = Variable.add_fixes_implicit @{term "x :: nat"} @{context}
 | 
|  |     82 | in
 | 
|  |     83 |   val _ =
 | 
|  |     84 |     map_range (fn n => (n + 1, mk (n + 1))) 20
 | 
|  |     85 |     |> map (fn (n, t) => 
 | 
|  |     86 |          Pretty.block [Pretty.str (Int.toString n ^ ": "),
 | 
|  |     87 |            Syntax.pretty_term ctxt' t] |> Pretty.writeln)
 | 
|  |     88 | end
 | 
|  |     89 | \<close>
 | 
|  |     90 | 
 | 
|  |     91 | (* test parsing *)
 | 
|  |     92 | term "Suc x"
 | 
|  |     93 | term "Suc\<^sup>2 x"
 | 
|  |     94 | term "Suc\<^sup>3 x"
 | 
|  |     95 | term "Suc\<^sup>4 x"
 | 
|  |     96 | term "Suc\<^sup>5 x"
 | 
|  |     97 | term "Suc\<^sup>6 x"
 | 
|  |     98 | term "Suc\<^sup>7 x"
 | 
|  |     99 | term "Suc\<^sup>8 x"
 | 
|  |    100 | term "Suc\<^sup>9 x"
 | 
|  |    101 | 
 | 
|  |    102 | term "Suc\<^bsup>2\<^esup> x"
 | 
|  |    103 | term "Suc\<^bsup>3\<^esup> x"
 | 
|  |    104 | term "Suc\<^bsup>4\<^esup> x"
 | 
|  |    105 | term "Suc\<^bsup>5\<^esup> x"
 | 
|  |    106 | term "Suc\<^bsup>6\<^esup> x"
 | 
|  |    107 | term "Suc\<^bsup>7\<^esup> x"
 | 
|  |    108 | term "Suc\<^bsup>8\<^esup> x"
 | 
|  |    109 | term "Suc\<^bsup>9\<^esup> x"
 | 
|  |    110 | term "Suc\<^bsup>10\<^esup> x"
 | 
|  |    111 | term "Suc\<^bsup>11\<^esup> x"
 | 
|  |    112 | term "Suc\<^bsup>12\<^esup> x"
 | 
|  |    113 | term "Suc\<^bsup>13\<^esup> x"
 | 
|  |    114 | term "Suc\<^bsup>14\<^esup> x"
 | 
|  |    115 | term "Suc\<^bsup>15\<^esup> x"
 | 
|  |    116 | term "Suc\<^bsup>16\<^esup> x"
 | 
|  |    117 | term "Suc\<^bsup>17\<^esup> x"
 | 
|  |    118 | term "Suc\<^bsup>18\<^esup> x"
 | 
|  |    119 | term "Suc\<^bsup>19\<^esup> x"
 | 
|  |    120 | term "Suc\<^bsup>20\<^esup> x"
 | 
|  |    121 | 
 | 
|  |    122 | (* check that the notation really means the right thing *)
 | 
|  |    123 | lemma "Suc\<^sup>2 n = n+2 \<and> Suc\<^sup>3 n = n+3 \<and> Suc\<^sup>4 n = n+4 \<and> Suc\<^sup>5 n = n+5
 | 
|  |    124 |   \<and> Suc\<^sup>6 n = n+6 \<and> Suc\<^sup>7 n = n+7 \<and> Suc\<^sup>8 n = n+8 \<and> Suc\<^sup>9 n = n+9"
 | 
|  |    125 | by simp
 | 
|  |    126 | 
 | 
|  |    127 | lemma "Suc\<^bsup>10\<^esup> n = n+10 \<and> Suc\<^bsup>11\<^esup> n = n+11 \<and> Suc\<^bsup>12\<^esup> n = n+12 \<and> Suc\<^bsup>13\<^esup> n = n+13
 | 
|  |    128 |   \<and> Suc\<^bsup>14\<^esup> n = n+14 \<and> Suc\<^bsup>15\<^esup> n = n+15 \<and> Suc\<^bsup>16\<^esup> n = n+16 \<and> Suc\<^bsup>17\<^esup> n = n+17
 | 
|  |    129 |   \<and> Suc\<^bsup>18\<^esup> n = n+18 \<and> Suc\<^bsup>19\<^esup> n = n+19 \<and> Suc\<^bsup>20\<^esup> n = n+20"
 | 
|  |    130 | by simp
 | 
|  |    131 | 
 | 
|  |    132 | *)
 | 
|  |    133 | 
 | 
|  |    134 | end |