src/HOL/NumberTheory/BijectionRel.thy
author nipkow
Tue, 08 Oct 2002 08:20:17 +0200
changeset 13630 a013a9dd370f
parent 13524 604d0f3622d6
child 16417 9bc16273c2d4
permissions -rw-r--r--
Got rid of rotates because of new simplifier
     1 (*  Title:      HOL/NumberTheory/BijectionRel.thy
     2     ID:         $Id$
     3     Author:     Thomas M. Rasmussen
     4     Copyright   2000  University of Cambridge
     5 *)
     6 
     7 header {* Bijections between sets *}
     8 
     9 theory BijectionRel = Main:
    10 
    11 text {*
    12   Inductive definitions of bijections between two different sets and
    13   between the same set.  Theorem for relating the two definitions.
    14 
    15   \bigskip
    16 *}
    17 
    18 consts
    19   bijR :: "('a => 'b => bool) => ('a set * 'b set) set"
    20 
    21 inductive "bijR P"
    22   intros
    23   empty [simp]: "({}, {}) \<in> bijR P"
    24   insert: "P a b ==> a \<notin> A ==> b \<notin> B ==> (A, B) \<in> bijR P
    25     ==> (insert a A, insert b B) \<in> bijR P"
    26 
    27 text {*
    28   Add extra condition to @{term insert}: @{term "\<forall>b \<in> B. \<not> P a b"}
    29   (and similar for @{term A}).
    30 *}
    31 
    32 constdefs
    33   bijP :: "('a => 'a => bool) => 'a set => bool"
    34   "bijP P F == \<forall>a b. a \<in> F \<and> P a b --> b \<in> F"
    35 
    36   uniqP :: "('a => 'a => bool) => bool"
    37   "uniqP P == \<forall>a b c d. P a b \<and> P c d --> (a = c) = (b = d)"
    38 
    39   symP :: "('a => 'a => bool) => bool"
    40   "symP P == \<forall>a b. P a b = P b a"
    41 
    42 consts
    43   bijER :: "('a => 'a => bool) => 'a set set"
    44 
    45 inductive "bijER P"
    46   intros
    47   empty [simp]: "{} \<in> bijER P"
    48   insert1: "P a a ==> a \<notin> A ==> A \<in> bijER P ==> insert a A \<in> bijER P"
    49   insert2: "P a b ==> a \<noteq> b ==> a \<notin> A ==> b \<notin> A ==> A \<in> bijER P
    50     ==> insert a (insert b A) \<in> bijER P"
    51 
    52 
    53 text {* \medskip @{term bijR} *}
    54 
    55 lemma fin_bijRl: "(A, B) \<in> bijR P ==> finite A"
    56   apply (erule bijR.induct)
    57   apply auto
    58   done
    59 
    60 lemma fin_bijRr: "(A, B) \<in> bijR P ==> finite B"
    61   apply (erule bijR.induct)
    62   apply auto
    63   done
    64 
    65 lemma aux_induct:
    66   "finite F ==> F \<subseteq> A ==> P {} ==>
    67     (!!F a. F \<subseteq> A ==> a \<in> A ==> a \<notin> F ==> P F ==> P (insert a F))
    68   ==> P F"
    69 proof -
    70   case rule_context
    71   assume major: "finite F"
    72     and subs: "F \<subseteq> A"
    73   show ?thesis
    74     apply (rule subs [THEN rev_mp])
    75     apply (rule major [THEN finite_induct])
    76      apply (blast intro: rule_context)+
    77     done
    78 qed
    79 
    80 lemma inj_func_bijR_aux1:
    81     "A \<subseteq> B ==> a \<notin> A ==> a \<in> B ==> inj_on f B ==> f a \<notin> f ` A"
    82   apply (unfold inj_on_def)
    83   apply auto
    84   done
    85 
    86 lemma inj_func_bijR_aux2:
    87   "\<forall>a. a \<in> A --> P a (f a) ==> inj_on f A ==> finite A ==> F <= A
    88     ==> (F, f ` F) \<in> bijR P"
    89   apply (rule_tac F = F and A = A in aux_induct)
    90      apply (rule finite_subset)
    91       apply auto
    92   apply (rule bijR.insert)
    93      apply (rule_tac [3] inj_func_bijR_aux1)
    94         apply auto
    95   done
    96 
    97 lemma inj_func_bijR:
    98   "\<forall>a. a \<in> A --> P a (f a) ==> inj_on f A ==> finite A
    99     ==> (A, f ` A) \<in> bijR P"
   100   apply (rule inj_func_bijR_aux2)
   101      apply auto
   102   done
   103 
   104 
   105 text {* \medskip @{term bijER} *}
   106 
   107 lemma fin_bijER: "A \<in> bijER P ==> finite A"
   108   apply (erule bijER.induct)
   109     apply auto
   110   done
   111 
   112 lemma aux1:
   113   "a \<notin> A ==> a \<notin> B ==> F \<subseteq> insert a A ==> F \<subseteq> insert a B ==> a \<in> F
   114     ==> \<exists>C. F = insert a C \<and> a \<notin> C \<and> C <= A \<and> C <= B"
   115   apply (rule_tac x = "F - {a}" in exI)
   116   apply auto
   117   done
   118 
   119 lemma aux2: "a \<noteq> b ==> a \<notin> A ==> b \<notin> B ==> a \<in> F ==> b \<in> F
   120     ==> F \<subseteq> insert a A ==> F \<subseteq> insert b B
   121     ==> \<exists>C. F = insert a (insert b C) \<and> a \<notin> C \<and> b \<notin> C \<and> C \<subseteq> A \<and> C \<subseteq> B"
   122   apply (rule_tac x = "F - {a, b}" in exI)
   123   apply auto
   124   done
   125 
   126 lemma aux_uniq: "uniqP P ==> P a b ==> P c d ==> (a = c) = (b = d)"
   127   apply (unfold uniqP_def)
   128   apply auto
   129   done
   130 
   131 lemma aux_sym: "symP P ==> P a b = P b a"
   132   apply (unfold symP_def)
   133   apply auto
   134   done
   135 
   136 lemma aux_in1:
   137     "uniqP P ==> b \<notin> C ==> P b b ==> bijP P (insert b C) ==> bijP P C"
   138   apply (unfold bijP_def)
   139   apply auto
   140   apply (subgoal_tac "b \<noteq> a")
   141    prefer 2
   142    apply clarify
   143   apply (simp add: aux_uniq)
   144   apply auto
   145   done
   146 
   147 lemma aux_in2:
   148   "symP P ==> uniqP P ==> a \<notin> C ==> b \<notin> C ==> a \<noteq> b ==> P a b
   149     ==> bijP P (insert a (insert b C)) ==> bijP P C"
   150   apply (unfold bijP_def)
   151   apply auto
   152   apply (subgoal_tac "aa \<noteq> a")
   153    prefer 2
   154    apply clarify
   155   apply (subgoal_tac "aa \<noteq> b")
   156    prefer 2
   157    apply clarify
   158   apply (simp add: aux_uniq)
   159   apply (subgoal_tac "ba \<noteq> a")
   160    apply auto
   161   apply (subgoal_tac "P a aa")
   162    prefer 2
   163    apply (simp add: aux_sym)
   164   apply (subgoal_tac "b = aa")
   165    apply (rule_tac [2] iffD1)
   166     apply (rule_tac [2] a = a and c = a and P = P in aux_uniq)
   167       apply auto
   168   done
   169 
   170 lemma aux_foo: "\<forall>a b. Q a \<and> P a b --> R b ==> P a b ==> Q a ==> R b"
   171   apply auto
   172   done
   173 
   174 lemma aux_bij: "bijP P F ==> symP P ==> P a b ==> (a \<in> F) = (b \<in> F)"
   175   apply (unfold bijP_def)
   176   apply (rule iffI)
   177   apply (erule_tac [!] aux_foo)
   178       apply simp_all
   179   apply (rule iffD2)
   180    apply (rule_tac P = P in aux_sym)
   181    apply simp_all
   182   done
   183 
   184 
   185 lemma aux_bijRER:
   186   "(A, B) \<in> bijR P ==> uniqP P ==> symP P
   187     ==> \<forall>F. bijP P F \<and> F \<subseteq> A \<and> F \<subseteq> B --> F \<in> bijER P"
   188   apply (erule bijR.induct)
   189    apply simp
   190   apply (case_tac "a = b")
   191    apply clarify
   192    apply (case_tac "b \<in> F")
   193     prefer 2
   194     apply (simp add: subset_insert)
   195    apply (cut_tac F = F and a = b and A = A and B = B in aux1)
   196         prefer 6
   197         apply clarify
   198         apply (rule bijER.insert1)
   199           apply simp_all
   200    apply (subgoal_tac "bijP P C")
   201     apply simp
   202    apply (rule aux_in1)
   203       apply simp_all
   204   apply clarify
   205   apply (case_tac "a \<in> F")
   206    apply (case_tac [!] "b \<in> F")
   207      apply (cut_tac F = F and a = a and b = b and A = A and B = B
   208        in aux2)
   209             apply (simp_all add: subset_insert)
   210     apply clarify
   211     apply (rule bijER.insert2)
   212         apply simp_all
   213     apply (subgoal_tac "bijP P C")
   214      apply simp
   215     apply (rule aux_in2)
   216           apply simp_all
   217    apply (subgoal_tac "b \<in> F")
   218     apply (rule_tac [2] iffD1)
   219      apply (rule_tac [2] a = a and F = F and P = P in aux_bij)
   220        apply (simp_all (no_asm_simp))
   221    apply (subgoal_tac [2] "a \<in> F")
   222     apply (rule_tac [3] iffD2)
   223      apply (rule_tac [3] b = b and F = F and P = P in aux_bij)
   224        apply auto
   225   done
   226 
   227 lemma bijR_bijER:
   228   "(A, A) \<in> bijR P ==>
   229     bijP P A ==> uniqP P ==> symP P ==> A \<in> bijER P"
   230   apply (cut_tac A = A and B = A and P = P in aux_bijRER)
   231      apply auto
   232   done
   233 
   234 end