src/HOL/Library/Finite_Cartesian_Product.thy
author wenzelm
Sun, 08 Mar 2009 00:16:34 +0100
changeset 30354 0d037d7e2a75
parent 30305 720226bedc4d
child 30488 5c4c3a9e9102
permissions -rw-r--r--
merged
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
29835
62da280e5d0b A formalization of finite cartesian product types
chaieb
parents:
diff changeset
     1
(* Title:      HOL/Library/Finite_Cartesian_Product
62da280e5d0b A formalization of finite cartesian product types
chaieb
parents:
diff changeset
     2
   Author:     Amine Chaieb, University of Cambridge
62da280e5d0b A formalization of finite cartesian product types
chaieb
parents:
diff changeset
     3
*)
62da280e5d0b A formalization of finite cartesian product types
chaieb
parents:
diff changeset
     4
62da280e5d0b A formalization of finite cartesian product types
chaieb
parents:
diff changeset
     5
header {* Definition of finite Cartesian product types. *}
62da280e5d0b A formalization of finite cartesian product types
chaieb
parents:
diff changeset
     6
62da280e5d0b A formalization of finite cartesian product types
chaieb
parents:
diff changeset
     7
theory Finite_Cartesian_Product
29841
86d94bb79226 A generic decision procedure for linear rea arithmetic and normed vector spaces
chaieb
parents: 29835
diff changeset
     8
  (* imports Plain SetInterval ATP_Linkup *)
86d94bb79226 A generic decision procedure for linear rea arithmetic and normed vector spaces
chaieb
parents: 29835
diff changeset
     9
imports Main
29835
62da280e5d0b A formalization of finite cartesian product types
chaieb
parents:
diff changeset
    10
begin
62da280e5d0b A formalization of finite cartesian product types
chaieb
parents:
diff changeset
    11
62da280e5d0b A formalization of finite cartesian product types
chaieb
parents:
diff changeset
    12
  (* FIXME : ATP_Linkup is only needed for metis at a few places. We could dispense of that by changing the proofs*)
62da280e5d0b A formalization of finite cartesian product types
chaieb
parents:
diff changeset
    13
subsection{* Dimention of sets *}
62da280e5d0b A formalization of finite cartesian product types
chaieb
parents:
diff changeset
    14
62da280e5d0b A formalization of finite cartesian product types
chaieb
parents:
diff changeset
    15
definition "dimindex (S:: 'a set) = (if finite (UNIV::'a set) then card (UNIV:: 'a set) else 1)"
62da280e5d0b A formalization of finite cartesian product types
chaieb
parents:
diff changeset
    16
62da280e5d0b A formalization of finite cartesian product types
chaieb
parents:
diff changeset
    17
syntax "_type_dimindex" :: "type => nat" ("(1DIM/(1'(_')))")
30304
d8e4cd2ac2a1 set operations Int, Un, INTER, UNION, Inter, Union, empty, UNIV are now proper qualified constants with authentic syntax
haftmann
parents: 29906
diff changeset
    18
translations "DIM(t)" => "CONST dimindex (CONST UNIV :: t set)"
29835
62da280e5d0b A formalization of finite cartesian product types
chaieb
parents:
diff changeset
    19
62da280e5d0b A formalization of finite cartesian product types
chaieb
parents:
diff changeset
    20
lemma dimindex_nonzero: "dimindex S \<noteq>  0"
62da280e5d0b A formalization of finite cartesian product types
chaieb
parents:
diff changeset
    21
unfolding dimindex_def 
62da280e5d0b A formalization of finite cartesian product types
chaieb
parents:
diff changeset
    22
by (simp add: neq0_conv[symmetric] del: neq0_conv)
62da280e5d0b A formalization of finite cartesian product types
chaieb
parents:
diff changeset
    23
62da280e5d0b A formalization of finite cartesian product types
chaieb
parents:
diff changeset
    24
lemma dimindex_ge_1: "dimindex S \<ge> 1"
62da280e5d0b A formalization of finite cartesian product types
chaieb
parents:
diff changeset
    25
  using dimindex_nonzero[of S] by arith 
62da280e5d0b A formalization of finite cartesian product types
chaieb
parents:
diff changeset
    26
lemma dimindex_univ: "dimindex (S :: 'a set) = DIM('a)" by (simp add: dimindex_def)
62da280e5d0b A formalization of finite cartesian product types
chaieb
parents:
diff changeset
    27
62da280e5d0b A formalization of finite cartesian product types
chaieb
parents:
diff changeset
    28
definition hassize (infixr "hassize" 12) where
62da280e5d0b A formalization of finite cartesian product types
chaieb
parents:
diff changeset
    29
  "(S hassize n) = (finite S \<and> card S = n)"
62da280e5d0b A formalization of finite cartesian product types
chaieb
parents:
diff changeset
    30
62da280e5d0b A formalization of finite cartesian product types
chaieb
parents:
diff changeset
    31
lemma dimindex_unique: " (UNIV :: 'a set) hassize n ==> DIM('a) = n"
62da280e5d0b A formalization of finite cartesian product types
chaieb
parents:
diff changeset
    32
by (simp add: dimindex_def hassize_def)
62da280e5d0b A formalization of finite cartesian product types
chaieb
parents:
diff changeset
    33
62da280e5d0b A formalization of finite cartesian product types
chaieb
parents:
diff changeset
    34
29906
80369da39838 section -> subsection
huffman
parents: 29841
diff changeset
    35
subsection{* An indexing type parametrized by base type. *}
29835
62da280e5d0b A formalization of finite cartesian product types
chaieb
parents:
diff changeset
    36
62da280e5d0b A formalization of finite cartesian product types
chaieb
parents:
diff changeset
    37
typedef 'a finite_image = "{1 .. DIM('a)}"
62da280e5d0b A formalization of finite cartesian product types
chaieb
parents:
diff changeset
    38
  using dimindex_ge_1 by auto
62da280e5d0b A formalization of finite cartesian product types
chaieb
parents:
diff changeset
    39
62da280e5d0b A formalization of finite cartesian product types
chaieb
parents:
diff changeset
    40
lemma finite_image_image: "(UNIV :: 'a finite_image set) = Abs_finite_image ` {1 .. DIM('a)}"
62da280e5d0b A formalization of finite cartesian product types
chaieb
parents:
diff changeset
    41
apply (auto simp add: Abs_finite_image_inverse image_def finite_image_def)
62da280e5d0b A formalization of finite cartesian product types
chaieb
parents:
diff changeset
    42
apply (rule_tac x="Rep_finite_image x" in bexI)
62da280e5d0b A formalization of finite cartesian product types
chaieb
parents:
diff changeset
    43
apply (simp_all add: Rep_finite_image_inverse Rep_finite_image)
62da280e5d0b A formalization of finite cartesian product types
chaieb
parents:
diff changeset
    44
using Rep_finite_image[where ?'a = 'a]
62da280e5d0b A formalization of finite cartesian product types
chaieb
parents:
diff changeset
    45
unfolding finite_image_def
62da280e5d0b A formalization of finite cartesian product types
chaieb
parents:
diff changeset
    46
apply simp
62da280e5d0b A formalization of finite cartesian product types
chaieb
parents:
diff changeset
    47
done
62da280e5d0b A formalization of finite cartesian product types
chaieb
parents:
diff changeset
    48
62da280e5d0b A formalization of finite cartesian product types
chaieb
parents:
diff changeset
    49
text{* Dimension of such a type, and indexing over it. *}
62da280e5d0b A formalization of finite cartesian product types
chaieb
parents:
diff changeset
    50
62da280e5d0b A formalization of finite cartesian product types
chaieb
parents:
diff changeset
    51
lemma inj_on_Abs_finite_image: 
62da280e5d0b A formalization of finite cartesian product types
chaieb
parents:
diff changeset
    52
  "inj_on (Abs_finite_image:: _ \<Rightarrow> 'a finite_image) {1 .. DIM('a)}"
62da280e5d0b A formalization of finite cartesian product types
chaieb
parents:
diff changeset
    53
by (auto simp add: inj_on_def finite_image_def Abs_finite_image_inject[where ?'a='a])
62da280e5d0b A formalization of finite cartesian product types
chaieb
parents:
diff changeset
    54
62da280e5d0b A formalization of finite cartesian product types
chaieb
parents:
diff changeset
    55
lemma has_size_finite_image: "(UNIV:: 'a finite_image set) hassize dimindex (S :: 'a set)"
62da280e5d0b A formalization of finite cartesian product types
chaieb
parents:
diff changeset
    56
  unfolding hassize_def finite_image_image card_image[OF inj_on_Abs_finite_image[where ?'a='a]] by (auto simp add: dimindex_def)
62da280e5d0b A formalization of finite cartesian product types
chaieb
parents:
diff changeset
    57
62da280e5d0b A formalization of finite cartesian product types
chaieb
parents:
diff changeset
    58
lemma hassize_image_inj: assumes f: "inj_on f S" and S: "S hassize n"
62da280e5d0b A formalization of finite cartesian product types
chaieb
parents:
diff changeset
    59
  shows "f ` S hassize n"
62da280e5d0b A formalization of finite cartesian product types
chaieb
parents:
diff changeset
    60
  using f S card_image[OF f]
62da280e5d0b A formalization of finite cartesian product types
chaieb
parents:
diff changeset
    61
    by (simp add: hassize_def inj_on_def)
62da280e5d0b A formalization of finite cartesian product types
chaieb
parents:
diff changeset
    62
62da280e5d0b A formalization of finite cartesian product types
chaieb
parents:
diff changeset
    63
lemma card_finite_image: "card (UNIV:: 'a finite_image set) = dimindex(S:: 'a set)"
62da280e5d0b A formalization of finite cartesian product types
chaieb
parents:
diff changeset
    64
using has_size_finite_image
62da280e5d0b A formalization of finite cartesian product types
chaieb
parents:
diff changeset
    65
unfolding hassize_def by blast
62da280e5d0b A formalization of finite cartesian product types
chaieb
parents:
diff changeset
    66
62da280e5d0b A formalization of finite cartesian product types
chaieb
parents:
diff changeset
    67
lemma finite_finite_image: "finite (UNIV:: 'a finite_image set)"
62da280e5d0b A formalization of finite cartesian product types
chaieb
parents:
diff changeset
    68
using has_size_finite_image
62da280e5d0b A formalization of finite cartesian product types
chaieb
parents:
diff changeset
    69
unfolding hassize_def by blast
62da280e5d0b A formalization of finite cartesian product types
chaieb
parents:
diff changeset
    70
62da280e5d0b A formalization of finite cartesian product types
chaieb
parents:
diff changeset
    71
lemma dimindex_finite_image: "dimindex (S:: 'a finite_image set) = dimindex(T:: 'a set)"
62da280e5d0b A formalization of finite cartesian product types
chaieb
parents:
diff changeset
    72
unfolding card_finite_image[of T, symmetric]
62da280e5d0b A formalization of finite cartesian product types
chaieb
parents:
diff changeset
    73
by (auto simp add: dimindex_def finite_finite_image)
62da280e5d0b A formalization of finite cartesian product types
chaieb
parents:
diff changeset
    74
62da280e5d0b A formalization of finite cartesian product types
chaieb
parents:
diff changeset
    75
lemma Abs_finite_image_works: 
62da280e5d0b A formalization of finite cartesian product types
chaieb
parents:
diff changeset
    76
  fixes i:: "'a finite_image"
62da280e5d0b A formalization of finite cartesian product types
chaieb
parents:
diff changeset
    77
  shows " \<exists>!n \<in> {1 .. DIM('a)}. Abs_finite_image n = i"
62da280e5d0b A formalization of finite cartesian product types
chaieb
parents:
diff changeset
    78
  unfolding Bex1_def Ex1_def
62da280e5d0b A formalization of finite cartesian product types
chaieb
parents:
diff changeset
    79
  apply (rule_tac x="Rep_finite_image i" in exI)
62da280e5d0b A formalization of finite cartesian product types
chaieb
parents:
diff changeset
    80
  using Rep_finite_image_inverse[where ?'a = 'a] 
62da280e5d0b A formalization of finite cartesian product types
chaieb
parents:
diff changeset
    81
    Rep_finite_image[where ?'a = 'a] 
62da280e5d0b A formalization of finite cartesian product types
chaieb
parents:
diff changeset
    82
  Abs_finite_image_inverse[where ?'a='a, symmetric]
62da280e5d0b A formalization of finite cartesian product types
chaieb
parents:
diff changeset
    83
  by (auto simp add: finite_image_def)
62da280e5d0b A formalization of finite cartesian product types
chaieb
parents:
diff changeset
    84
62da280e5d0b A formalization of finite cartesian product types
chaieb
parents:
diff changeset
    85
lemma Abs_finite_image_inj: 
62da280e5d0b A formalization of finite cartesian product types
chaieb
parents:
diff changeset
    86
 "i \<in> {1 .. DIM('a)} \<Longrightarrow> j \<in> {1 .. DIM('a)}
62da280e5d0b A formalization of finite cartesian product types
chaieb
parents:
diff changeset
    87
  \<Longrightarrow> (((Abs_finite_image i ::'a finite_image) = Abs_finite_image j) \<longleftrightarrow> (i = j))"
62da280e5d0b A formalization of finite cartesian product types
chaieb
parents:
diff changeset
    88
  using Abs_finite_image_works[where ?'a = 'a] 
62da280e5d0b A formalization of finite cartesian product types
chaieb
parents:
diff changeset
    89
  by (auto simp add: atLeastAtMost_iff Bex1_def)
62da280e5d0b A formalization of finite cartesian product types
chaieb
parents:
diff changeset
    90
62da280e5d0b A formalization of finite cartesian product types
chaieb
parents:
diff changeset
    91
lemma forall_Abs_finite_image: 
62da280e5d0b A formalization of finite cartesian product types
chaieb
parents:
diff changeset
    92
  "(\<forall>k:: 'a finite_image. P k) \<longleftrightarrow> (\<forall>i \<in> {1 .. DIM('a)}. P(Abs_finite_image i))"
62da280e5d0b A formalization of finite cartesian product types
chaieb
parents:
diff changeset
    93
unfolding Ball_def atLeastAtMost_iff Ex1_def
62da280e5d0b A formalization of finite cartesian product types
chaieb
parents:
diff changeset
    94
using Abs_finite_image_works[where ?'a = 'a, unfolded atLeastAtMost_iff Bex1_def]
62da280e5d0b A formalization of finite cartesian product types
chaieb
parents:
diff changeset
    95
by metis
62da280e5d0b A formalization of finite cartesian product types
chaieb
parents:
diff changeset
    96
62da280e5d0b A formalization of finite cartesian product types
chaieb
parents:
diff changeset
    97
subsection {* Finite Cartesian products, with indexing and lambdas. *}
62da280e5d0b A formalization of finite cartesian product types
chaieb
parents:
diff changeset
    98
62da280e5d0b A formalization of finite cartesian product types
chaieb
parents:
diff changeset
    99
typedef (Cart)
62da280e5d0b A formalization of finite cartesian product types
chaieb
parents:
diff changeset
   100
  ('a, 'b) "^" (infixl "^" 15)
62da280e5d0b A formalization of finite cartesian product types
chaieb
parents:
diff changeset
   101
    = "{f:: 'b finite_image \<Rightarrow> 'a . True}" by simp
62da280e5d0b A formalization of finite cartesian product types
chaieb
parents:
diff changeset
   102
62da280e5d0b A formalization of finite cartesian product types
chaieb
parents:
diff changeset
   103
abbreviation dimset:: "('a ^ 'n) \<Rightarrow> nat set" where
62da280e5d0b A formalization of finite cartesian product types
chaieb
parents:
diff changeset
   104
  "dimset a \<equiv> {1 .. DIM('n)}"
62da280e5d0b A formalization of finite cartesian product types
chaieb
parents:
diff changeset
   105
62da280e5d0b A formalization of finite cartesian product types
chaieb
parents:
diff changeset
   106
definition Cart_nth :: "'a ^ 'b \<Rightarrow> nat \<Rightarrow> 'a" (infixl "$" 90) where
62da280e5d0b A formalization of finite cartesian product types
chaieb
parents:
diff changeset
   107
  "x$i = Rep_Cart x (Abs_finite_image i)"
62da280e5d0b A formalization of finite cartesian product types
chaieb
parents:
diff changeset
   108
62da280e5d0b A formalization of finite cartesian product types
chaieb
parents:
diff changeset
   109
lemma stupid_ext: "(\<forall>x. f x = g x) \<longleftrightarrow> (f = g)"
62da280e5d0b A formalization of finite cartesian product types
chaieb
parents:
diff changeset
   110
  apply auto
62da280e5d0b A formalization of finite cartesian product types
chaieb
parents:
diff changeset
   111
  apply (rule ext)
62da280e5d0b A formalization of finite cartesian product types
chaieb
parents:
diff changeset
   112
  apply auto
62da280e5d0b A formalization of finite cartesian product types
chaieb
parents:
diff changeset
   113
  done
62da280e5d0b A formalization of finite cartesian product types
chaieb
parents:
diff changeset
   114
lemma Cart_eq: "((x:: 'a ^ 'b) = y) \<longleftrightarrow> (\<forall>i\<in> dimset x. x$i = y$i)"
62da280e5d0b A formalization of finite cartesian product types
chaieb
parents:
diff changeset
   115
  unfolding Cart_nth_def forall_Abs_finite_image[symmetric, where P = "\<lambda>i. Rep_Cart x i = Rep_Cart y i"] stupid_ext
62da280e5d0b A formalization of finite cartesian product types
chaieb
parents:
diff changeset
   116
  using Rep_Cart_inject[of x y] ..
62da280e5d0b A formalization of finite cartesian product types
chaieb
parents:
diff changeset
   117
62da280e5d0b A formalization of finite cartesian product types
chaieb
parents:
diff changeset
   118
consts Cart_lambda :: "(nat \<Rightarrow> 'a) \<Rightarrow> 'a ^ 'b" 
62da280e5d0b A formalization of finite cartesian product types
chaieb
parents:
diff changeset
   119
notation (xsymbols) Cart_lambda (binder "\<chi>" 10)
62da280e5d0b A formalization of finite cartesian product types
chaieb
parents:
diff changeset
   120
62da280e5d0b A formalization of finite cartesian product types
chaieb
parents:
diff changeset
   121
defs Cart_lambda_def: "Cart_lambda g == (SOME (f:: 'a ^ 'b). \<forall>i \<in> {1 .. DIM('b)}. f$i = g i)"
62da280e5d0b A formalization of finite cartesian product types
chaieb
parents:
diff changeset
   122
62da280e5d0b A formalization of finite cartesian product types
chaieb
parents:
diff changeset
   123
lemma  Cart_lambda_beta: " \<forall> i\<in> {1 .. DIM('b)}. (Cart_lambda g:: 'a ^ 'b)$i = g i"
62da280e5d0b A formalization of finite cartesian product types
chaieb
parents:
diff changeset
   124
  unfolding Cart_lambda_def
62da280e5d0b A formalization of finite cartesian product types
chaieb
parents:
diff changeset
   125
proof (rule someI_ex)
62da280e5d0b A formalization of finite cartesian product types
chaieb
parents:
diff changeset
   126
  let ?p = "\<lambda>(i::nat) (k::'b finite_image). i \<in> {1 .. DIM('b)} \<and> (Abs_finite_image i = k)"
62da280e5d0b A formalization of finite cartesian product types
chaieb
parents:
diff changeset
   127
  let ?f = "Abs_Cart (\<lambda>k. g (THE i. ?p i k)):: 'a ^ 'b"
62da280e5d0b A formalization of finite cartesian product types
chaieb
parents:
diff changeset
   128
  let ?P = "\<lambda>f i. f$i = g i"
62da280e5d0b A formalization of finite cartesian product types
chaieb
parents:
diff changeset
   129
  let ?Q = "\<lambda>(f::'a ^ 'b). \<forall> i \<in> {1 .. DIM('b)}. ?P f i"
62da280e5d0b A formalization of finite cartesian product types
chaieb
parents:
diff changeset
   130
  {fix i 
62da280e5d0b A formalization of finite cartesian product types
chaieb
parents:
diff changeset
   131
    assume i: "i \<in> {1 .. DIM('b)}"
62da280e5d0b A formalization of finite cartesian product types
chaieb
parents:
diff changeset
   132
    let ?j = "THE j. ?p j (Abs_finite_image i)"
62da280e5d0b A formalization of finite cartesian product types
chaieb
parents:
diff changeset
   133
    from theI'[where P = "\<lambda>j. ?p (j::nat) (Abs_finite_image i :: 'b finite_image)", OF Abs_finite_image_works[of "Abs_finite_image i :: 'b finite_image", unfolded Bex1_def]]
62da280e5d0b A formalization of finite cartesian product types
chaieb
parents:
diff changeset
   134
    have j: "?j \<in> {1 .. DIM('b)}" "(Abs_finite_image ?j :: 'b finite_image) = Abs_finite_image i" by blast+
62da280e5d0b A formalization of finite cartesian product types
chaieb
parents:
diff changeset
   135
    from i j Abs_finite_image_inject[of i ?j, where ?'a = 'b]
62da280e5d0b A formalization of finite cartesian product types
chaieb
parents:
diff changeset
   136
    have th: "?j = i" by (simp add: finite_image_def)  
62da280e5d0b A formalization of finite cartesian product types
chaieb
parents:
diff changeset
   137
    have "?P ?f i"
62da280e5d0b A formalization of finite cartesian product types
chaieb
parents:
diff changeset
   138
      using th
62da280e5d0b A formalization of finite cartesian product types
chaieb
parents:
diff changeset
   139
      by (simp add: Cart_nth_def Abs_Cart_inverse Rep_Cart_inverse Cart_def) }
62da280e5d0b A formalization of finite cartesian product types
chaieb
parents:
diff changeset
   140
  hence th0: "?Q ?f" ..
62da280e5d0b A formalization of finite cartesian product types
chaieb
parents:
diff changeset
   141
  with th0 show "\<exists>f. ?Q f" unfolding Ex1_def by auto
62da280e5d0b A formalization of finite cartesian product types
chaieb
parents:
diff changeset
   142
qed
62da280e5d0b A formalization of finite cartesian product types
chaieb
parents:
diff changeset
   143
62da280e5d0b A formalization of finite cartesian product types
chaieb
parents:
diff changeset
   144
lemma  Cart_lambda_beta': "i\<in> {1 .. DIM('b)} \<Longrightarrow> (Cart_lambda g:: 'a ^ 'b)$i = g i"
62da280e5d0b A formalization of finite cartesian product types
chaieb
parents:
diff changeset
   145
  using Cart_lambda_beta by blast
62da280e5d0b A formalization of finite cartesian product types
chaieb
parents:
diff changeset
   146
62da280e5d0b A formalization of finite cartesian product types
chaieb
parents:
diff changeset
   147
lemma Cart_lambda_unique:
62da280e5d0b A formalization of finite cartesian product types
chaieb
parents:
diff changeset
   148
  fixes f :: "'a ^ 'b"
62da280e5d0b A formalization of finite cartesian product types
chaieb
parents:
diff changeset
   149
  shows "(\<forall>i\<in> {1 .. DIM('b)}. f$i = g i) \<longleftrightarrow> Cart_lambda g = f"
62da280e5d0b A formalization of finite cartesian product types
chaieb
parents:
diff changeset
   150
  by (auto simp add: Cart_eq Cart_lambda_beta)
62da280e5d0b A formalization of finite cartesian product types
chaieb
parents:
diff changeset
   151
62da280e5d0b A formalization of finite cartesian product types
chaieb
parents:
diff changeset
   152
lemma Cart_lambda_eta: "(\<chi> i. (g$i)) = g" by (simp add: Cart_eq Cart_lambda_beta)
62da280e5d0b A formalization of finite cartesian product types
chaieb
parents:
diff changeset
   153
62da280e5d0b A formalization of finite cartesian product types
chaieb
parents:
diff changeset
   154
text{* A non-standard sum to "paste" Cartesian products. *}
62da280e5d0b A formalization of finite cartesian product types
chaieb
parents:
diff changeset
   155
62da280e5d0b A formalization of finite cartesian product types
chaieb
parents:
diff changeset
   156
typedef ('a,'b) finite_sum = "{1 .. DIM('a) + DIM('b)}"
62da280e5d0b A formalization of finite cartesian product types
chaieb
parents:
diff changeset
   157
  apply (rule exI[where x="1"])
62da280e5d0b A formalization of finite cartesian product types
chaieb
parents:
diff changeset
   158
  using dimindex_ge_1[of "UNIV :: 'a set"] dimindex_ge_1[of "UNIV :: 'b set"]
62da280e5d0b A formalization of finite cartesian product types
chaieb
parents:
diff changeset
   159
  by auto
62da280e5d0b A formalization of finite cartesian product types
chaieb
parents:
diff changeset
   160
62da280e5d0b A formalization of finite cartesian product types
chaieb
parents:
diff changeset
   161
definition pastecart :: "'a ^ 'm \<Rightarrow> 'a ^ 'n \<Rightarrow> 'a ^ ('m,'n) finite_sum" where
62da280e5d0b A formalization of finite cartesian product types
chaieb
parents:
diff changeset
   162
  "pastecart f g = (\<chi> i. (if i <= DIM('m) then f$i else g$(i - DIM('m))))"
62da280e5d0b A formalization of finite cartesian product types
chaieb
parents:
diff changeset
   163
62da280e5d0b A formalization of finite cartesian product types
chaieb
parents:
diff changeset
   164
definition fstcart:: "'a ^('m, 'n) finite_sum \<Rightarrow> 'a ^ 'm" where
62da280e5d0b A formalization of finite cartesian product types
chaieb
parents:
diff changeset
   165
  "fstcart f = (\<chi> i. (f$i))"
62da280e5d0b A formalization of finite cartesian product types
chaieb
parents:
diff changeset
   166
62da280e5d0b A formalization of finite cartesian product types
chaieb
parents:
diff changeset
   167
definition sndcart:: "'a ^('m, 'n) finite_sum \<Rightarrow> 'a ^ 'n" where
62da280e5d0b A formalization of finite cartesian product types
chaieb
parents:
diff changeset
   168
  "sndcart f = (\<chi> i. (f$(i + DIM('m))))"
62da280e5d0b A formalization of finite cartesian product types
chaieb
parents:
diff changeset
   169
62da280e5d0b A formalization of finite cartesian product types
chaieb
parents:
diff changeset
   170
lemma finite_sum_image: "(UNIV::('a,'b) finite_sum set) = Abs_finite_sum ` {1 .. DIM('a) + DIM('b)}"
62da280e5d0b A formalization of finite cartesian product types
chaieb
parents:
diff changeset
   171
apply (auto  simp add: image_def)
62da280e5d0b A formalization of finite cartesian product types
chaieb
parents:
diff changeset
   172
apply (rule_tac x="Rep_finite_sum x" in bexI)
62da280e5d0b A formalization of finite cartesian product types
chaieb
parents:
diff changeset
   173
apply (simp add: Rep_finite_sum_inverse)
62da280e5d0b A formalization of finite cartesian product types
chaieb
parents:
diff changeset
   174
using Rep_finite_sum[unfolded finite_sum_def, where ?'a = 'a and ?'b = 'b]
62da280e5d0b A formalization of finite cartesian product types
chaieb
parents:
diff changeset
   175
apply (simp add: Rep_finite_sum)
62da280e5d0b A formalization of finite cartesian product types
chaieb
parents:
diff changeset
   176
done
62da280e5d0b A formalization of finite cartesian product types
chaieb
parents:
diff changeset
   177
62da280e5d0b A formalization of finite cartesian product types
chaieb
parents:
diff changeset
   178
lemma inj_on_Abs_finite_sum: "inj_on (Abs_finite_sum :: _ \<Rightarrow> ('a,'b) finite_sum) {1 .. DIM('a) + DIM('b)}" 
62da280e5d0b A formalization of finite cartesian product types
chaieb
parents:
diff changeset
   179
  using Abs_finite_sum_inject[where ?'a = 'a and ?'b = 'b]
62da280e5d0b A formalization of finite cartesian product types
chaieb
parents:
diff changeset
   180
  by (auto simp add: inj_on_def finite_sum_def)
62da280e5d0b A formalization of finite cartesian product types
chaieb
parents:
diff changeset
   181
62da280e5d0b A formalization of finite cartesian product types
chaieb
parents:
diff changeset
   182
lemma dimindex_has_size_finite_sum:
62da280e5d0b A formalization of finite cartesian product types
chaieb
parents:
diff changeset
   183
  "(UNIV::('m,'n) finite_sum set) hassize (DIM('m) + DIM('n))"
62da280e5d0b A formalization of finite cartesian product types
chaieb
parents:
diff changeset
   184
  by (simp add: finite_sum_image hassize_def card_image[OF inj_on_Abs_finite_sum[where ?'a = 'm and ?'b = 'n]] del: One_nat_def)
62da280e5d0b A formalization of finite cartesian product types
chaieb
parents:
diff changeset
   185
62da280e5d0b A formalization of finite cartesian product types
chaieb
parents:
diff changeset
   186
lemma dimindex_finite_sum: "DIM(('m,'n) finite_sum) = DIM('m) + DIM('n)"
62da280e5d0b A formalization of finite cartesian product types
chaieb
parents:
diff changeset
   187
  using dimindex_has_size_finite_sum[where ?'n = 'n and ?'m = 'm, unfolded hassize_def]
62da280e5d0b A formalization of finite cartesian product types
chaieb
parents:
diff changeset
   188
  by (simp add: dimindex_def)
62da280e5d0b A formalization of finite cartesian product types
chaieb
parents:
diff changeset
   189
62da280e5d0b A formalization of finite cartesian product types
chaieb
parents:
diff changeset
   190
lemma fstcart_pastecart: "fstcart (pastecart (x::'a ^'m ) (y:: 'a ^ 'n)) = x"
62da280e5d0b A formalization of finite cartesian product types
chaieb
parents:
diff changeset
   191
  by (simp add: pastecart_def fstcart_def Cart_eq Cart_lambda_beta dimindex_finite_sum)
62da280e5d0b A formalization of finite cartesian product types
chaieb
parents:
diff changeset
   192
62da280e5d0b A formalization of finite cartesian product types
chaieb
parents:
diff changeset
   193
lemma sndcart_pastecart: "sndcart (pastecart (x::'a ^'m ) (y:: 'a ^ 'n)) = y"
62da280e5d0b A formalization of finite cartesian product types
chaieb
parents:
diff changeset
   194
  by (simp add: pastecart_def sndcart_def Cart_eq Cart_lambda_beta dimindex_finite_sum)
62da280e5d0b A formalization of finite cartesian product types
chaieb
parents:
diff changeset
   195
62da280e5d0b A formalization of finite cartesian product types
chaieb
parents:
diff changeset
   196
lemma pastecart_fst_snd: "pastecart (fstcart z) (sndcart z) = z"
62da280e5d0b A formalization of finite cartesian product types
chaieb
parents:
diff changeset
   197
proof -
62da280e5d0b A formalization of finite cartesian product types
chaieb
parents:
diff changeset
   198
 {fix i
62da280e5d0b A formalization of finite cartesian product types
chaieb
parents:
diff changeset
   199
  assume H: "i \<le> DIM('b) + DIM('c)" 
62da280e5d0b A formalization of finite cartesian product types
chaieb
parents:
diff changeset
   200
    "\<not> i \<le> DIM('b)"
62da280e5d0b A formalization of finite cartesian product types
chaieb
parents:
diff changeset
   201
    from H have ith: "i - DIM('b) \<in> {1 .. DIM('c)}"
62da280e5d0b A formalization of finite cartesian product types
chaieb
parents:
diff changeset
   202
      apply simp by arith
62da280e5d0b A formalization of finite cartesian product types
chaieb
parents:
diff changeset
   203
    from H have th0: "i - DIM('b) + DIM('b) = i"
62da280e5d0b A formalization of finite cartesian product types
chaieb
parents:
diff changeset
   204
      by simp
62da280e5d0b A formalization of finite cartesian product types
chaieb
parents:
diff changeset
   205
  have "(\<chi> i. (z$(i + DIM('b))) :: 'a ^ 'c)$(i - DIM('b)) = z$i"
62da280e5d0b A formalization of finite cartesian product types
chaieb
parents:
diff changeset
   206
    unfolding Cart_lambda_beta'[where g = "\<lambda> i. z$(i + DIM('b))", OF ith] th0 ..}
62da280e5d0b A formalization of finite cartesian product types
chaieb
parents:
diff changeset
   207
thus ?thesis by (auto simp add: pastecart_def fstcart_def sndcart_def Cart_eq Cart_lambda_beta dimindex_finite_sum)
62da280e5d0b A formalization of finite cartesian product types
chaieb
parents:
diff changeset
   208
qed
62da280e5d0b A formalization of finite cartesian product types
chaieb
parents:
diff changeset
   209
62da280e5d0b A formalization of finite cartesian product types
chaieb
parents:
diff changeset
   210
lemma pastecart_eq: "(x = y) \<longleftrightarrow> (fstcart x = fstcart y) \<and> (sndcart x = sndcart y)"
62da280e5d0b A formalization of finite cartesian product types
chaieb
parents:
diff changeset
   211
  using pastecart_fst_snd[of x] pastecart_fst_snd[of y] by metis
62da280e5d0b A formalization of finite cartesian product types
chaieb
parents:
diff changeset
   212
62da280e5d0b A formalization of finite cartesian product types
chaieb
parents:
diff changeset
   213
lemma forall_pastecart: "(\<forall>p. P p) \<longleftrightarrow> (\<forall>x y. P (pastecart x y))"
62da280e5d0b A formalization of finite cartesian product types
chaieb
parents:
diff changeset
   214
  by (metis pastecart_fst_snd fstcart_pastecart sndcart_pastecart)
62da280e5d0b A formalization of finite cartesian product types
chaieb
parents:
diff changeset
   215
62da280e5d0b A formalization of finite cartesian product types
chaieb
parents:
diff changeset
   216
lemma exists_pastecart: "(\<exists>p. P p)  \<longleftrightarrow> (\<exists>x y. P (pastecart x y))"
62da280e5d0b A formalization of finite cartesian product types
chaieb
parents:
diff changeset
   217
  by (metis pastecart_fst_snd fstcart_pastecart sndcart_pastecart)
62da280e5d0b A formalization of finite cartesian product types
chaieb
parents:
diff changeset
   218
62da280e5d0b A formalization of finite cartesian product types
chaieb
parents:
diff changeset
   219
text{* The finiteness lemma. *}
62da280e5d0b A formalization of finite cartesian product types
chaieb
parents:
diff changeset
   220
62da280e5d0b A formalization of finite cartesian product types
chaieb
parents:
diff changeset
   221
lemma finite_cart:
62da280e5d0b A formalization of finite cartesian product types
chaieb
parents:
diff changeset
   222
 "\<forall>i \<in> {1 .. DIM('n)}. finite {x.  P i x}
62da280e5d0b A formalization of finite cartesian product types
chaieb
parents:
diff changeset
   223
  \<Longrightarrow> finite {v::'a ^ 'n . (\<forall>i \<in> {1 .. DIM('n)}. P i (v$i))}"
62da280e5d0b A formalization of finite cartesian product types
chaieb
parents:
diff changeset
   224
proof-
62da280e5d0b A formalization of finite cartesian product types
chaieb
parents:
diff changeset
   225
  assume f: "\<forall>i \<in> {1 .. DIM('n)}. finite {x.  P i x}"
62da280e5d0b A formalization of finite cartesian product types
chaieb
parents:
diff changeset
   226
  {fix n
62da280e5d0b A formalization of finite cartesian product types
chaieb
parents:
diff changeset
   227
    assume n: "n \<le> DIM('n)"
62da280e5d0b A formalization of finite cartesian product types
chaieb
parents:
diff changeset
   228
    have "finite {v:: 'a ^ 'n . (\<forall>i\<in> {1 .. DIM('n)}. i \<le> n \<longrightarrow> P i (v$i))
62da280e5d0b A formalization of finite cartesian product types
chaieb
parents:
diff changeset
   229
                              \<and> (\<forall>i\<in> {1 .. DIM('n)}. n < i \<longrightarrow> v$i = (SOME x. False))}" 
62da280e5d0b A formalization of finite cartesian product types
chaieb
parents:
diff changeset
   230
      using n 
62da280e5d0b A formalization of finite cartesian product types
chaieb
parents:
diff changeset
   231
      proof(induct n)
62da280e5d0b A formalization of finite cartesian product types
chaieb
parents:
diff changeset
   232
	case 0
62da280e5d0b A formalization of finite cartesian product types
chaieb
parents:
diff changeset
   233
	have th0: "{v . (\<forall>i \<in> {1 .. DIM('n)}. v$i = (SOME x. False))} =
62da280e5d0b A formalization of finite cartesian product types
chaieb
parents:
diff changeset
   234
      {(\<chi> i. (SOME x. False)::'a ^ 'n)}" by (auto simp add: Cart_lambda_beta Cart_eq)
62da280e5d0b A formalization of finite cartesian product types
chaieb
parents:
diff changeset
   235
	with "0.prems" show ?case by auto
62da280e5d0b A formalization of finite cartesian product types
chaieb
parents:
diff changeset
   236
      next
62da280e5d0b A formalization of finite cartesian product types
chaieb
parents:
diff changeset
   237
	case (Suc n)
62da280e5d0b A formalization of finite cartesian product types
chaieb
parents:
diff changeset
   238
	let ?h = "\<lambda>(x::'a,v:: 'a ^ 'n). (\<chi> i. if i = Suc n then x else v$i):: 'a ^ 'n"
62da280e5d0b A formalization of finite cartesian product types
chaieb
parents:
diff changeset
   239
	let ?T = "{v\<Colon>'a ^ 'n.
62da280e5d0b A formalization of finite cartesian product types
chaieb
parents:
diff changeset
   240
            (\<forall>i\<Colon>nat\<in>{1\<Colon>nat..DIM('n)}. i \<le> Suc n \<longrightarrow> P i (v$i)) \<and>
62da280e5d0b A formalization of finite cartesian product types
chaieb
parents:
diff changeset
   241
            (\<forall>i\<Colon>nat\<in>{1\<Colon>nat..DIM('n)}.
62da280e5d0b A formalization of finite cartesian product types
chaieb
parents:
diff changeset
   242
                Suc n < i \<longrightarrow> v$i = (SOME x\<Colon>'a. False))}"
62da280e5d0b A formalization of finite cartesian product types
chaieb
parents:
diff changeset
   243
	let ?S = "{x::'a . P (Suc  n) x} \<times> {v:: 'a^'n. (\<forall>i \<in> {1 .. DIM('n)}. i <= n \<longrightarrow> P i (v$i)) \<and> (\<forall>i \<in> {1 .. DIM('n)}. n < i \<longrightarrow> v$i = (SOME x. False))}"
62da280e5d0b A formalization of finite cartesian product types
chaieb
parents:
diff changeset
   244
	have th0: " ?T \<subseteq> (?h ` ?S)" 
62da280e5d0b A formalization of finite cartesian product types
chaieb
parents:
diff changeset
   245
	  using Suc.prems
62da280e5d0b A formalization of finite cartesian product types
chaieb
parents:
diff changeset
   246
	  apply (auto simp add: image_def)
62da280e5d0b A formalization of finite cartesian product types
chaieb
parents:
diff changeset
   247
	  apply (rule_tac x = "x$(Suc n)" in exI)
62da280e5d0b A formalization of finite cartesian product types
chaieb
parents:
diff changeset
   248
	  apply (rule conjI)
62da280e5d0b A formalization of finite cartesian product types
chaieb
parents:
diff changeset
   249
	  apply (rotate_tac)
62da280e5d0b A formalization of finite cartesian product types
chaieb
parents:
diff changeset
   250
	  apply (erule ballE[where x="Suc n"])
62da280e5d0b A formalization of finite cartesian product types
chaieb
parents:
diff changeset
   251
	  apply simp
62da280e5d0b A formalization of finite cartesian product types
chaieb
parents:
diff changeset
   252
	  apply simp
62da280e5d0b A formalization of finite cartesian product types
chaieb
parents:
diff changeset
   253
	  apply (rule_tac x= "\<chi> i. if i = Suc n then (SOME x:: 'a. False) else (x:: 'a ^ 'n)$i:: 'a ^ 'n" in exI)
62da280e5d0b A formalization of finite cartesian product types
chaieb
parents:
diff changeset
   254
	  by (simp add: Cart_eq Cart_lambda_beta)
62da280e5d0b A formalization of finite cartesian product types
chaieb
parents:
diff changeset
   255
	have th1: "finite ?S" 
62da280e5d0b A formalization of finite cartesian product types
chaieb
parents:
diff changeset
   256
	  apply (rule finite_cartesian_product) 
62da280e5d0b A formalization of finite cartesian product types
chaieb
parents:
diff changeset
   257
	  using f Suc.hyps Suc.prems by auto 
62da280e5d0b A formalization of finite cartesian product types
chaieb
parents:
diff changeset
   258
	from finite_imageI[OF th1] have th2: "finite (?h ` ?S)" . 
62da280e5d0b A formalization of finite cartesian product types
chaieb
parents:
diff changeset
   259
	from finite_subset[OF th0 th2] show ?case by blast 
62da280e5d0b A formalization of finite cartesian product types
chaieb
parents:
diff changeset
   260
      qed}
62da280e5d0b A formalization of finite cartesian product types
chaieb
parents:
diff changeset
   261
62da280e5d0b A formalization of finite cartesian product types
chaieb
parents:
diff changeset
   262
  note th = this
62da280e5d0b A formalization of finite cartesian product types
chaieb
parents:
diff changeset
   263
  from this[of "DIM('n)"] f
62da280e5d0b A formalization of finite cartesian product types
chaieb
parents:
diff changeset
   264
  show ?thesis by auto
62da280e5d0b A formalization of finite cartesian product types
chaieb
parents:
diff changeset
   265
qed
62da280e5d0b A formalization of finite cartesian product types
chaieb
parents:
diff changeset
   266
62da280e5d0b A formalization of finite cartesian product types
chaieb
parents:
diff changeset
   267
62da280e5d0b A formalization of finite cartesian product types
chaieb
parents:
diff changeset
   268
end