section ‹Bijections between sets›
theory BijectionRel
imports Main
begin
text ‹
Inductive definitions of bijections between two different sets and
between the same set. Theorem for relating the two definitions.
\bigskip
›
inductive_set
bijR :: "('a => 'b => bool) => ('a set * 'b set) set"
for P :: "'a => 'b => bool"
where
empty [simp]: "({}, {}) ∈ bijR P"
| insert: "P a b ==> a ∉ A ==> b ∉ B ==> (A, B) ∈ bijR P
==> (insert a A, insert b B) ∈ bijR P"
text ‹
Add extra condition to @{term insert}: @{term "∀b ∈ B. ¬ P a b"}
(and similar for @{term A}).
›
definition
bijP :: "('a => 'a => bool) => 'a set => bool" where
"bijP P F = (∀a b. a ∈ F ∧ P a b --> b ∈ F)"
definition
uniqP :: "('a => 'a => bool) => bool" where
"uniqP P = (∀a b c d. P a b ∧ P c d --> (a = c) = (b = d))"
definition
symP :: "('a => 'a => bool) => bool" where
"symP P = (∀a b. P a b = P b a)"
inductive_set
bijER :: "('a => 'a => bool) => 'a set set"
for P :: "'a => 'a => bool"
where
empty [simp]: "{} ∈ bijER P"
| insert1: "P a a ==> a ∉ A ==> A ∈ bijER P ==> insert a A ∈ bijER P"
| insert2: "P a b ==> a ≠ b ==> a ∉ A ==> b ∉ A ==> A ∈ bijER P
==> insert a (insert b A) ∈ bijER P"
text ‹\medskip @{term bijR}›
lemma fin_bijRl: "(A, B) ∈ bijR P ==> finite A"
apply (erule bijR.induct)
apply auto
done
lemma fin_bijRr: "(A, B) ∈ bijR P ==> finite B"
apply (erule bijR.induct)
apply auto
done
lemma aux_induct:
assumes major: "finite F"
and subs: "F ⊆ A"
and cases: "P {}"
"!!F a. F ⊆ A ==> a ∈ A ==> a ∉ F ==> P F ==> P (insert a F)"
shows "P F"
using major subs
apply (induct set: finite)
apply (blast intro: cases)+
done
lemma inj_func_bijR_aux1:
"A ⊆ B ==> a ∉ A ==> a ∈ B ==> inj_on f B ==> f a ∉ f ` A"
apply (unfold inj_on_def)
apply auto
done
lemma inj_func_bijR_aux2:
"∀a. a ∈ A --> P a (f a) ==> inj_on f A ==> finite A ==> F <= A
==> (F, f ` F) ∈ bijR P"
apply (rule_tac F = F and A = A in aux_induct)
apply (rule finite_subset)
apply auto
apply (rule bijR.insert)
apply (rule_tac [3] inj_func_bijR_aux1)
apply auto
done
lemma inj_func_bijR:
"∀a. a ∈ A --> P a (f a) ==> inj_on f A ==> finite A
==> (A, f ` A) ∈ bijR P"
apply (rule inj_func_bijR_aux2)
apply auto
done
text ‹\medskip @{term bijER}›
lemma fin_bijER: "A ∈ bijER P ==> finite A"
apply (erule bijER.induct)
apply auto
done
lemma aux1:
"a ∉ A ==> a ∉ B ==> F ⊆ insert a A ==> F ⊆ insert a B ==> a ∈ F
==> ∃C. F = insert a C ∧ a ∉ C ∧ C <= A ∧ C <= B"
apply (rule_tac x = "F - {a}" in exI)
apply auto
done
lemma aux2: "a ≠ b ==> a ∉ A ==> b ∉ B ==> a ∈ F ==> b ∈ F
==> F ⊆ insert a A ==> F ⊆ insert b B
==> ∃C. F = insert a (insert b C) ∧ a ∉ C ∧ b ∉ C ∧ C ⊆ A ∧ C ⊆ B"
apply (rule_tac x = "F - {a, b}" in exI)
apply auto
done
lemma aux_uniq: "uniqP P ==> P a b ==> P c d ==> (a = c) = (b = d)"
apply (unfold uniqP_def)
apply auto
done
lemma aux_sym: "symP P ==> P a b = P b a"
apply (unfold symP_def)
apply auto
done
lemma aux_in1:
"uniqP P ==> b ∉ C ==> P b b ==> bijP P (insert b C) ==> bijP P C"
apply (unfold bijP_def)
apply auto
apply (subgoal_tac "b ≠ a")
prefer 2
apply clarify
apply (simp add: aux_uniq)
apply auto
done
lemma aux_in2:
"symP P ==> uniqP P ==> a ∉ C ==> b ∉ C ==> a ≠ b ==> P a b
==> bijP P (insert a (insert b C)) ==> bijP P C"
apply (unfold bijP_def)
apply auto
apply (subgoal_tac "aa ≠ a")
prefer 2
apply clarify
apply (subgoal_tac "aa ≠ b")
prefer 2
apply clarify
apply (simp add: aux_uniq)
apply (subgoal_tac "ba ≠ a")
apply auto
apply (subgoal_tac "P a aa")
prefer 2
apply (simp add: aux_sym)
apply (subgoal_tac "b = aa")
apply (rule_tac [2] iffD1)
apply (rule_tac [2] a = a and c = a and P = P in aux_uniq)
apply auto
done
lemma aux_foo: "∀a b. Q a ∧ P a b --> R b ==> P a b ==> Q a ==> R b"
apply auto
done
lemma aux_bij: "bijP P F ==> symP P ==> P a b ==> (a ∈ F) = (b ∈ F)"
apply (unfold bijP_def)
apply (rule iffI)
apply (erule_tac [!] aux_foo)
apply simp_all
apply (rule iffD2)
apply (rule_tac P = P in aux_sym)
apply simp_all
done
lemma aux_bijRER:
"(A, B) ∈ bijR P ==> uniqP P ==> symP P
==> ∀F. bijP P F ∧ F ⊆ A ∧ F ⊆ B --> F ∈ bijER P"
apply (erule bijR.induct)
apply simp
apply (case_tac "a = b")
apply clarify
apply (case_tac "b ∈ F")
prefer 2
apply (simp add: subset_insert)
apply (cut_tac F = F and a = b and A = A and B = B in aux1)
prefer 6
apply clarify
apply (rule bijER.insert1)
apply simp_all
apply (subgoal_tac "bijP P C")
apply simp
apply (rule aux_in1)
apply simp_all
apply clarify
apply (case_tac "a ∈ F")
apply (case_tac [!] "b ∈ F")
apply (cut_tac F = F and a = a and b = b and A = A and B = B
in aux2)
apply (simp_all add: subset_insert)
apply clarify
apply (rule bijER.insert2)
apply simp_all
apply (subgoal_tac "bijP P C")
apply simp
apply (rule aux_in2)
apply simp_all
apply (subgoal_tac "b ∈ F")
apply (rule_tac [2] iffD1)
apply (rule_tac [2] a = a and F = F and P = P in aux_bij)
apply (simp_all (no_asm_simp))
apply (subgoal_tac [2] "a ∈ F")
apply (rule_tac [3] iffD2)
apply (rule_tac [3] b = b and F = F and P = P in aux_bij)
apply auto
done
lemma bijR_bijER:
"(A, A) ∈ bijR P ==>
bijP P A ==> uniqP P ==> symP P ==> A ∈ bijER P"
apply (cut_tac A = A and B = A and P = P in aux_bijRER)
apply auto
done
end