src/HOL/Library/Type_Length.thy
author wenzelm
Mon, 30 Nov 2020 17:10:49 +0100
changeset 72792 26492b600d78
parent 72281 beeadb35e357
permissions -rw-r--r--
tuned whitespace --- avoid TABs;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
63762
6920b1885eff clarified session;
wenzelm
parents: 61799
diff changeset
     1
(*  Title:      HOL/Library/Type_Length.thy
6920b1885eff clarified session;
wenzelm
parents: 61799
diff changeset
     2
    Author:     John Matthews, Galois Connections, Inc., Copyright 2006
24465
70f0214b3ecc revert to Word library version from 2007/08/20
huffman
parents:
diff changeset
     3
*)
25262
d0928156e326 Added reference to Jeremy Dawson's paper on the word library.
kleing
parents: 24465
diff changeset
     4
63762
6920b1885eff clarified session;
wenzelm
parents: 61799
diff changeset
     5
section \<open>Assigning lengths to types by type classes\<close>
25262
d0928156e326 Added reference to Jeremy Dawson's paper on the word library.
kleing
parents: 24465
diff changeset
     6
37655
f4d616d41a59 more speaking theory names
haftmann
parents: 29631
diff changeset
     7
theory Type_Length
63762
6920b1885eff clarified session;
wenzelm
parents: 61799
diff changeset
     8
imports Numeral_Type
24465
70f0214b3ecc revert to Word library version from 2007/08/20
huffman
parents:
diff changeset
     9
begin
70f0214b3ecc revert to Word library version from 2007/08/20
huffman
parents:
diff changeset
    10
61799
4cf66f21b764 isabelle update_cartouches -c -t;
wenzelm
parents: 58874
diff changeset
    11
text \<open>
24465
70f0214b3ecc revert to Word library version from 2007/08/20
huffman
parents:
diff changeset
    12
  The aim of this is to allow any type as index type, but to provide a
70f0214b3ecc revert to Word library version from 2007/08/20
huffman
parents:
diff changeset
    13
  default instantiation for numeral types. This independence requires
63762
6920b1885eff clarified session;
wenzelm
parents: 61799
diff changeset
    14
  some duplication with the definitions in \<^file>\<open>Numeral_Type.thy\<close>.
61799
4cf66f21b764 isabelle update_cartouches -c -t;
wenzelm
parents: 58874
diff changeset
    15
\<close>
24465
70f0214b3ecc revert to Word library version from 2007/08/20
huffman
parents:
diff changeset
    16
29608
564ea783ace8 no base sort in class import
haftmann
parents: 26560
diff changeset
    17
class len0 =
26514
eff55c0a6d34 tuned towards code generation
haftmann
parents: 25262
diff changeset
    18
  fixes len_of :: "'a itself \<Rightarrow> nat"
24465
70f0214b3ecc revert to Word library version from 2007/08/20
huffman
parents:
diff changeset
    19
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69593
diff changeset
    20
syntax "_type_length" :: "type \<Rightarrow> nat" (\<open>(1LENGTH/(1'(_')))\<close>)
64113
86efd3d4dc98 dedicated syntax for types with a length
haftmann
parents: 63762
diff changeset
    21
86efd3d4dc98 dedicated syntax for types with a length
haftmann
parents: 63762
diff changeset
    22
translations "LENGTH('a)" \<rightharpoonup>
86efd3d4dc98 dedicated syntax for types with a length
haftmann
parents: 63762
diff changeset
    23
  "CONST len_of (CONST Pure.type :: 'a itself)"
86efd3d4dc98 dedicated syntax for types with a length
haftmann
parents: 63762
diff changeset
    24
86efd3d4dc98 dedicated syntax for types with a length
haftmann
parents: 63762
diff changeset
    25
print_translation \<open>
86efd3d4dc98 dedicated syntax for types with a length
haftmann
parents: 63762
diff changeset
    26
  let
69593
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 64113
diff changeset
    27
    fun len_of_itself_tr' ctxt [Const (\<^const_syntax>\<open>Pure.type\<close>, Type (_, [T]))] =
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 64113
diff changeset
    28
      Syntax.const \<^syntax_const>\<open>_type_length\<close> $ Syntax_Phases.term_of_typ ctxt T
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 64113
diff changeset
    29
  in [(\<^const_syntax>\<open>len_of\<close>, len_of_itself_tr')] end
64113
86efd3d4dc98 dedicated syntax for types with a length
haftmann
parents: 63762
diff changeset
    30
\<close>
86efd3d4dc98 dedicated syntax for types with a length
haftmann
parents: 63762
diff changeset
    31
63762
6920b1885eff clarified session;
wenzelm
parents: 61799
diff changeset
    32
text \<open>Some theorems are only true on words with length greater 0.\<close>
26514
eff55c0a6d34 tuned towards code generation
haftmann
parents: 25262
diff changeset
    33
eff55c0a6d34 tuned towards code generation
haftmann
parents: 25262
diff changeset
    34
class len = len0 +
64113
86efd3d4dc98 dedicated syntax for types with a length
haftmann
parents: 63762
diff changeset
    35
  assumes len_gt_0 [iff]: "0 < LENGTH('a)"
70348
bde161c740ca more theorems for proof of concept for word type
haftmann
parents: 69597
diff changeset
    36
begin
bde161c740ca more theorems for proof of concept for word type
haftmann
parents: 69597
diff changeset
    37
bde161c740ca more theorems for proof of concept for word type
haftmann
parents: 69597
diff changeset
    38
lemma len_not_eq_0 [simp]:
bde161c740ca more theorems for proof of concept for word type
haftmann
parents: 69597
diff changeset
    39
  "LENGTH('a) \<noteq> 0"
bde161c740ca more theorems for proof of concept for word type
haftmann
parents: 69597
diff changeset
    40
  by simp
bde161c740ca more theorems for proof of concept for word type
haftmann
parents: 69597
diff changeset
    41
bde161c740ca more theorems for proof of concept for word type
haftmann
parents: 69597
diff changeset
    42
end
26514
eff55c0a6d34 tuned towards code generation
haftmann
parents: 25262
diff changeset
    43
eff55c0a6d34 tuned towards code generation
haftmann
parents: 25262
diff changeset
    44
instantiation num0 and num1 :: len0
eff55c0a6d34 tuned towards code generation
haftmann
parents: 25262
diff changeset
    45
begin
eff55c0a6d34 tuned towards code generation
haftmann
parents: 25262
diff changeset
    46
63762
6920b1885eff clarified session;
wenzelm
parents: 61799
diff changeset
    47
definition len_num0: "len_of (_ :: num0 itself) = 0"
6920b1885eff clarified session;
wenzelm
parents: 61799
diff changeset
    48
definition len_num1: "len_of (_ :: num1 itself) = 1"
26514
eff55c0a6d34 tuned towards code generation
haftmann
parents: 25262
diff changeset
    49
eff55c0a6d34 tuned towards code generation
haftmann
parents: 25262
diff changeset
    50
instance ..
24465
70f0214b3ecc revert to Word library version from 2007/08/20
huffman
parents:
diff changeset
    51
26514
eff55c0a6d34 tuned towards code generation
haftmann
parents: 25262
diff changeset
    52
end
eff55c0a6d34 tuned towards code generation
haftmann
parents: 25262
diff changeset
    53
eff55c0a6d34 tuned towards code generation
haftmann
parents: 25262
diff changeset
    54
instantiation bit0 and bit1 :: (len0) len0
eff55c0a6d34 tuned towards code generation
haftmann
parents: 25262
diff changeset
    55
begin
24465
70f0214b3ecc revert to Word library version from 2007/08/20
huffman
parents:
diff changeset
    56
64113
86efd3d4dc98 dedicated syntax for types with a length
haftmann
parents: 63762
diff changeset
    57
definition len_bit0: "len_of (_ :: 'a::len0 bit0 itself) = 2 * LENGTH('a)"
86efd3d4dc98 dedicated syntax for types with a length
haftmann
parents: 63762
diff changeset
    58
definition len_bit1: "len_of (_ :: 'a::len0 bit1 itself) = 2 * LENGTH('a) + 1"
26514
eff55c0a6d34 tuned towards code generation
haftmann
parents: 25262
diff changeset
    59
eff55c0a6d34 tuned towards code generation
haftmann
parents: 25262
diff changeset
    60
instance ..
eff55c0a6d34 tuned towards code generation
haftmann
parents: 25262
diff changeset
    61
eff55c0a6d34 tuned towards code generation
haftmann
parents: 25262
diff changeset
    62
end
24465
70f0214b3ecc revert to Word library version from 2007/08/20
huffman
parents:
diff changeset
    63
70f0214b3ecc revert to Word library version from 2007/08/20
huffman
parents:
diff changeset
    64
lemmas len_of_numeral_defs [simp] = len_num0 len_num1 len_bit0 len_bit1
70f0214b3ecc revert to Word library version from 2007/08/20
huffman
parents:
diff changeset
    65
63762
6920b1885eff clarified session;
wenzelm
parents: 61799
diff changeset
    66
instance num1 :: len
6920b1885eff clarified session;
wenzelm
parents: 61799
diff changeset
    67
  by standard simp
6920b1885eff clarified session;
wenzelm
parents: 61799
diff changeset
    68
instance bit0 :: (len) len
6920b1885eff clarified session;
wenzelm
parents: 61799
diff changeset
    69
  by standard simp
6920b1885eff clarified session;
wenzelm
parents: 61799
diff changeset
    70
instance bit1 :: (len0) len
6920b1885eff clarified session;
wenzelm
parents: 61799
diff changeset
    71
  by standard simp
24465
70f0214b3ecc revert to Word library version from 2007/08/20
huffman
parents:
diff changeset
    72
70901
94a0c47b8553 moved quickcheck setup to distribution
haftmann
parents: 70348
diff changeset
    73
instantiation Enum.finite_1 :: len
94a0c47b8553 moved quickcheck setup to distribution
haftmann
parents: 70348
diff changeset
    74
begin
94a0c47b8553 moved quickcheck setup to distribution
haftmann
parents: 70348
diff changeset
    75
94a0c47b8553 moved quickcheck setup to distribution
haftmann
parents: 70348
diff changeset
    76
definition
94a0c47b8553 moved quickcheck setup to distribution
haftmann
parents: 70348
diff changeset
    77
  "len_of_finite_1 (x :: Enum.finite_1 itself) \<equiv> (1 :: nat)"
94a0c47b8553 moved quickcheck setup to distribution
haftmann
parents: 70348
diff changeset
    78
94a0c47b8553 moved quickcheck setup to distribution
haftmann
parents: 70348
diff changeset
    79
instance
94a0c47b8553 moved quickcheck setup to distribution
haftmann
parents: 70348
diff changeset
    80
  by standard (auto simp: len_of_finite_1_def)
94a0c47b8553 moved quickcheck setup to distribution
haftmann
parents: 70348
diff changeset
    81
24465
70f0214b3ecc revert to Word library version from 2007/08/20
huffman
parents:
diff changeset
    82
end
70901
94a0c47b8553 moved quickcheck setup to distribution
haftmann
parents: 70348
diff changeset
    83
94a0c47b8553 moved quickcheck setup to distribution
haftmann
parents: 70348
diff changeset
    84
instantiation Enum.finite_2 :: len
94a0c47b8553 moved quickcheck setup to distribution
haftmann
parents: 70348
diff changeset
    85
begin
94a0c47b8553 moved quickcheck setup to distribution
haftmann
parents: 70348
diff changeset
    86
94a0c47b8553 moved quickcheck setup to distribution
haftmann
parents: 70348
diff changeset
    87
definition
94a0c47b8553 moved quickcheck setup to distribution
haftmann
parents: 70348
diff changeset
    88
  "len_of_finite_2 (x :: Enum.finite_2 itself) \<equiv> (2 :: nat)"
94a0c47b8553 moved quickcheck setup to distribution
haftmann
parents: 70348
diff changeset
    89
94a0c47b8553 moved quickcheck setup to distribution
haftmann
parents: 70348
diff changeset
    90
instance
94a0c47b8553 moved quickcheck setup to distribution
haftmann
parents: 70348
diff changeset
    91
  by standard (auto simp: len_of_finite_2_def)
94a0c47b8553 moved quickcheck setup to distribution
haftmann
parents: 70348
diff changeset
    92
94a0c47b8553 moved quickcheck setup to distribution
haftmann
parents: 70348
diff changeset
    93
end
94a0c47b8553 moved quickcheck setup to distribution
haftmann
parents: 70348
diff changeset
    94
94a0c47b8553 moved quickcheck setup to distribution
haftmann
parents: 70348
diff changeset
    95
instantiation Enum.finite_3 :: len
94a0c47b8553 moved quickcheck setup to distribution
haftmann
parents: 70348
diff changeset
    96
begin
94a0c47b8553 moved quickcheck setup to distribution
haftmann
parents: 70348
diff changeset
    97
94a0c47b8553 moved quickcheck setup to distribution
haftmann
parents: 70348
diff changeset
    98
definition
94a0c47b8553 moved quickcheck setup to distribution
haftmann
parents: 70348
diff changeset
    99
  "len_of_finite_3 (x :: Enum.finite_3 itself) \<equiv> (4 :: nat)"
94a0c47b8553 moved quickcheck setup to distribution
haftmann
parents: 70348
diff changeset
   100
94a0c47b8553 moved quickcheck setup to distribution
haftmann
parents: 70348
diff changeset
   101
instance
94a0c47b8553 moved quickcheck setup to distribution
haftmann
parents: 70348
diff changeset
   102
  by standard (auto simp: len_of_finite_3_def)
94a0c47b8553 moved quickcheck setup to distribution
haftmann
parents: 70348
diff changeset
   103
94a0c47b8553 moved quickcheck setup to distribution
haftmann
parents: 70348
diff changeset
   104
end
94a0c47b8553 moved quickcheck setup to distribution
haftmann
parents: 70348
diff changeset
   105
71195
d50a718ccf35 tuned material
haftmann
parents: 70901
diff changeset
   106
lemma length_not_greater_eq_2_iff [simp]:
d50a718ccf35 tuned material
haftmann
parents: 70901
diff changeset
   107
  \<open>\<not> 2 \<le> LENGTH('a::len) \<longleftrightarrow> LENGTH('a) = 1\<close>
d50a718ccf35 tuned material
haftmann
parents: 70901
diff changeset
   108
  by (auto simp add: not_le dest: less_2_cases)
d50a718ccf35 tuned material
haftmann
parents: 70901
diff changeset
   109
71759
816e52bbfa60 more theorems
haftmann
parents: 71195
diff changeset
   110
context linordered_idom
816e52bbfa60 more theorems
haftmann
parents: 71195
diff changeset
   111
begin
816e52bbfa60 more theorems
haftmann
parents: 71195
diff changeset
   112
816e52bbfa60 more theorems
haftmann
parents: 71195
diff changeset
   113
lemma two_less_eq_exp_length [simp]:
816e52bbfa60 more theorems
haftmann
parents: 71195
diff changeset
   114
  \<open>2 \<le> 2 ^ LENGTH('b::len)\<close>
816e52bbfa60 more theorems
haftmann
parents: 71195
diff changeset
   115
  using mult_left_mono [of 1 \<open>2 ^ (LENGTH('b::len) - 1)\<close> 2]
72792
26492b600d78 tuned whitespace --- avoid TABs;
wenzelm
parents: 72281
diff changeset
   116
  by (cases \<open>LENGTH('b::len)\<close>) simp_all
71759
816e52bbfa60 more theorems
haftmann
parents: 71195
diff changeset
   117
70901
94a0c47b8553 moved quickcheck setup to distribution
haftmann
parents: 70348
diff changeset
   118
end
71759
816e52bbfa60 more theorems
haftmann
parents: 71195
diff changeset
   119
72281
beeadb35e357 more thorough treatment of division, particularly signed division on int and word
haftmann
parents: 71759
diff changeset
   120
lemma less_eq_decr_length_iff [simp]:
beeadb35e357 more thorough treatment of division, particularly signed division on int and word
haftmann
parents: 71759
diff changeset
   121
  \<open>n \<le> LENGTH('a::len) - Suc 0 \<longleftrightarrow> n < LENGTH('a)\<close>
beeadb35e357 more thorough treatment of division, particularly signed division on int and word
haftmann
parents: 71759
diff changeset
   122
  by (cases \<open>LENGTH('a)\<close>) (simp_all add: less_Suc_eq le_less)
beeadb35e357 more thorough treatment of division, particularly signed division on int and word
haftmann
parents: 71759
diff changeset
   123
beeadb35e357 more thorough treatment of division, particularly signed division on int and word
haftmann
parents: 71759
diff changeset
   124
lemma decr_length_less_iff [simp]:
beeadb35e357 more thorough treatment of division, particularly signed division on int and word
haftmann
parents: 71759
diff changeset
   125
  \<open>LENGTH('a::len) - Suc 0 < n \<longleftrightarrow> LENGTH('a) \<le> n\<close>
beeadb35e357 more thorough treatment of division, particularly signed division on int and word
haftmann
parents: 71759
diff changeset
   126
  by (cases \<open>LENGTH('a)\<close>) auto
beeadb35e357 more thorough treatment of division, particularly signed division on int and word
haftmann
parents: 71759
diff changeset
   127
71759
816e52bbfa60 more theorems
haftmann
parents: 71195
diff changeset
   128
end