Theory Matrix
section ‹Matrices›
theory Matrix
imports "HOL-Library.Word" Dioid
begin
text ‹In this section we formalise a perhaps more natural version of
matrices of fixed dimension ($m \times n$-matrices). It is well known
that such matrices over a Kleene algebra form a Kleene
algebra~\<^cite>‹"conway71regular"›.›
subsection ‹Type Definition›
typedef (overloaded) 'a atMost = "{..<LENGTH('a::len)}"
by auto
declare Rep_atMost_inject [simp]
lemma UNIV_atMost:
"(UNIV::'a atMost set) = Abs_atMost ` {..<LENGTH('a::len)}"
apply auto
apply (rule Abs_atMost_induct)
apply auto
done
lemma finite_UNIV_atMost [simp]: "finite (UNIV::('a::len) atMost set)"
by (simp add: UNIV_atMost)
text ‹Our matrix type is similar to \mbox{‹'a^'n^'m›} from
{\em HOL/Multivariate\_Analysis/Finite\_Cartesian\_Product.thy}, but
(i)~we explicitly define a type constructor for matrices and square
matrices, and (ii)~in the definition of operations, e.g., matrix
multiplication, we impose weaker sort requirements on the element
type.›
context notes [[typedef_overloaded]]
begin