Theory Basis
subsection ‹Definitions extending HOL as logical basis of Bali›
theory Basis
imports Main
begin
subsubsection "misc"
ML ‹fun strip_tac ctxt i = REPEAT (resolve_tac ctxt [impI, allI] i)›
declare if_split_asm [split] option.split [split] option.split_asm [split]
setup ‹map_theory_simpset (fn ctxt => ctxt addloop ("split_all_tac", split_all_tac))›
declare if_weak_cong [cong del] option.case_cong_weak [cong del]
declare length_Suc_conv [iff]
lemma Collect_split_eq: "{p. P (case_prod f p)} = {(a,b). P (f a b)}"
by auto
lemma subset_insertD: "A ⊆ insert x B ⟹ A ⊆ B ∧ x ∉ A ∨ (∃B'. A = insert x B' ∧ B' ⊆ B)"
apply (case_tac "x ∈ A")
apply (rule disjI2)
apply (rule_tac x = "A - {x}" in exI)
apply fast+
done
abbreviation nat3 :: nat ("3") where "3 ≡ Suc 2"
abbreviation nat4 :: nat ("4") where "4 ≡ Suc 3"
lemma irrefl_tranclI': "r¯ ∩ r⇧+ = {} ⟹ ∀x. (x, x) ∉ r⇧+"
by (blast elim: tranclE dest: trancl_into_rtrancl)
lemma trancl_rtrancl_trancl: "⟦(x, y) ∈ r⇧+; (y, z) ∈ r⇧*⟧ ⟹ (x, z) ∈ r⇧+"
by (auto dest: tranclD rtrancl_trans rtrancl_into_trancl2)
lemma rtrancl_into_trancl3: "⟦(a, b) ∈ r⇧*; a ≠ b⟧ ⟹ (a, b) ∈ r⇧+"
apply (drule rtranclD)
apply auto
done
lemma rtrancl_into_rtrancl2: "⟦(a, b) ∈ r; (b, c) ∈ r⇧*⟧ ⟹ (a, c) ∈ r⇧*"
by (auto intro: rtrancl_trans)
lemma triangle_lemma:
assumes unique: "⋀a b c. ⟦(a,b)∈r; (a,c)∈r⟧ ⟹ b = c"
and ax: "(a,x)∈r⇧*" and ay: "(a,y)∈r⇧*"
shows "(x,y)∈r⇧* ∨ (y,x)∈r⇧*"
using ax ay
proof (induct rule: converse_rtrancl_induct)
assume "(x,y)∈r⇧*"
then show ?thesis by blast
next
fix a v
assume a_v_r: "(a, v) ∈ r"
and v_x_rt: "(v, x) ∈ r⇧*"
and a_y_rt: "(a, y) ∈ r⇧*"
and hyp: "(v, y) ∈ r⇧* ⟹ (x, y) ∈ r⇧* ∨ (y, x) ∈ r⇧*"
from a_y_rt show "(x, y) ∈ r⇧* ∨ (y, x) ∈ r⇧*"
proof (cases rule: converse_rtranclE)
assume "a = y"
with a_v_r v_x_rt have "(y,x) ∈ r⇧*"
by (auto intro: rtrancl_trans)
then show ?thesis by blast
next
fix w
assume a_w_r: "(a, w) ∈ r"
and w_y_rt: "(w, y) ∈ r⇧*"
from a_v_r a_w_r unique have "v=w" by auto
with w_y_rt hyp show ?thesis by blast
qed
qed
lemma rtrancl_cases:
assumes "(a,b)∈r⇧*"
obtains (Refl) "a = b"
| (Trancl) "(a,b)∈r⇧+"
apply (rule rtranclE [OF assms])
apply (auto dest: rtrancl_into_trancl1)
done
lemma Ball_weaken: "⟦Ball s P; ⋀ x. P x⟶Q x⟧⟹Ball s Q"
by auto
lemma finite_SetCompr2:
"finite {f y x |x y. P y}" if "finite (Collect P)"
"∀y. P y ⟶ finite (range (f y))"
proof -
have "{f y x |x y. P y} = (⋃y∈Collect P. range (f y))"
by auto
with that show ?thesis by simp
qed
lemma list_all2_trans: "∀a b c. P1 a b ⟶ P2 b c ⟶ P3 a c ⟹
∀xs2 xs3. list_all2 P1 xs1 xs2 ⟶ list_all2 P2 xs2 xs3 ⟶ list_all2 P3 xs1 xs3"
apply (induct_tac xs1)
apply simp
apply (rule allI)
apply (induct_tac xs2)
apply simp
apply (rule allI)
apply (induct_tac xs3)
apply auto
done
subsubsection "pairs"
lemma surjective_pairing5:
"p = (fst p, fst (snd p), fst (snd (snd p)), fst (snd (snd (snd p))),
snd (snd (snd (snd p))))"
by auto
lemma fst_splitE [elim!]:
assumes "fst s' = x'"
obtains x s where "s' = (x,s)" and "x = x'"
using assms by (cases s') auto
lemma fst_in_set_lemma: "(x, y) ∈ set l ⟹ x ∈ fst ` set l"
by (induct l) auto
subsubsection "quantifiers"
lemma All_Ex_refl_eq2 [simp]: "(∀x. (∃b. x = f b ∧ Q b) ⟶ P x) = (∀b. Q b ⟶ P (f b))"
by auto
lemma ex_ex_miniscope1 [simp]: "(∃w v. P w v ∧ Q v) = (∃v. (∃w. P w v) ∧ Q v)"
by auto
lemma ex_miniscope2 [simp]: "(∃v. P v ∧ Q ∧ R v) = (Q ∧ (∃v. P v ∧ R v))"
by auto
lemma ex_reorder31: "(∃z x y. P x y z) = (∃x y z. P x y z)"
by auto
lemma All_Ex_refl_eq1 [simp]: "(∀x. (∃b. x = f b) ⟶ P x) = (∀b. P (f b))"
by auto
subsubsection "sums"
notation case_sum (infixr "'(+')" 80)
primrec the_Inl :: "'a + 'b ⇒ 'a"
where "the_Inl (Inl a) = a"
primrec the_Inr :: "'a + 'b ⇒ 'b"
where "the_Inr (Inr b) = b"