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