Theory LList
theory LList imports ZF begin
consts
llist :: "i=>i"
codatatype
"llist(A)" = LNil | LCons ("a ∈ A", "l ∈ llist(A)")
consts
lleq :: "i=>i"
coinductive
domains "lleq(A)" ⊆ "llist(A) * llist(A)"
intros
LNil: "<LNil, LNil> ∈ lleq(A)"
LCons: "[| a ∈ A; <l,l'> ∈ lleq(A) |]
==> <LCons(a,l), LCons(a,l')> ∈ lleq(A)"
type_intros llist.intros
definition
lconst :: "i => i" where
"lconst(a) == lfp(univ(a), %l. LCons(a,l))"
axiomatization flip :: "i => i"
where
flip_LNil: "flip(LNil) = LNil" and
flip_LCons: "[| x ∈ bool; l ∈ llist(bool) |]
==> flip(LCons(x,l)) = LCons(not(x), flip(l))"
declare subsetI [rule del]
subsetCE [rule del]
declare subset_refl [intro!]
cons_subsetI [intro!]
subset_consI [intro!]
Union_least [intro!]
UN_least [intro!]
Un_least [intro!]
Inter_greatest [intro!]
Int_greatest [intro!]
RepFun_subset [intro!]
Un_upper1 [intro!]
Un_upper2 [intro!]
Int_lower1 [intro!]
Int_lower2 [intro!]
inductive_cases LConsE: "LCons(a,l) ∈ llist(A)"
lemma LCons_iff: "LCons(a,l)=LCons(a',l') ⟷ a=a' & l=l'"
by auto
lemma LNil_LCons_iff: "~ LNil=LCons(a,l)"
by auto
lemma llist_mono: "A ⊆ B ==> llist(A) ⊆ llist(B)"
apply (unfold llist.defs )
apply (rule gfp_mono)
apply (rule llist.bnd_mono)
apply (assumption | rule quniv_mono basic_monos)+
done
declare QPair_Int_Vset_subset_UN [THEN subset_trans, intro!]
QPair_subset_univ [intro!]
empty_subsetI [intro!]
one_in_quniv [THEN qunivD, intro!]
declare qunivD [dest!]
declare Ord_in_Ord [elim!]
lemma llist_quniv_lemma:
"Ord(i) ==> l ∈ llist(quniv(A)) ⟹ l ∩ Vset(i) ⊆ univ(eclose(A))"
proof (induct i arbitrary: l rule: trans_induct)
case (step i l)
show ?case using ‹l ∈ llist(quniv(A))›
proof (cases l rule: llist.cases)
case LNil thus ?thesis
by (simp add: QInl_def QInr_def llist.con_defs)
next
case (LCons a l) thus ?thesis using step.hyps
proof (simp add: QInl_def QInr_def llist.con_defs)
show "<1; <a; l>> ∩ Vset(i) ⊆ univ(eclose(A))" using LCons ‹Ord(i)›
by (fast intro: step Ord_trans Int_lower1 [THEN subset_trans])
qed
qed
qed
lemma llist_quniv: "llist(quniv(A)) ⊆ quniv(A)"
apply (rule qunivI [THEN subsetI])
apply (rule Int_Vset_subset)
apply (assumption | rule llist_quniv_lemma)+
done
lemmas llist_subset_quniv =
subset_trans [OF llist_mono llist_quniv]
declare QPair_Int_Vset_subset_UN [THEN subset_trans, intro!]
QPair_mono [intro!]
declare Ord_in_Ord [elim!]
lemma lleq_Int_Vset_subset:
"Ord(i) ==> <l,l'> ∈ lleq(A) ⟹ l ∩ Vset(i) ⊆ l'"
proof (induct i arbitrary: l l' rule: trans_induct)
case (step i l l')
show ?case using ‹⟨l, l'⟩ ∈ lleq(A)›
proof (cases rule: lleq.cases)
case LNil thus ?thesis
by (auto simp add: QInr_def llist.con_defs)
next
case (LCons a l l') thus ?thesis using step.hyps
by (auto simp add: QInr_def llist.con_defs intro: Ord_trans)
qed
qed
lemma lleq_symmetric: "<l,l'> ∈ lleq(A) ==> <l',l> ∈ lleq(A)"
apply (erule lleq.coinduct [OF converseI])
apply (rule lleq.dom_subset [THEN converse_type], safe)
apply (erule lleq.cases, blast+)
done
lemma lleq_implies_equal: "<l,l'> ∈ lleq(A) ==> l=l'"
apply (rule equalityI)
apply (assumption | rule lleq_Int_Vset_subset [THEN Int_Vset_subset] |
erule lleq_symmetric)+
done
lemma equal_llist_implies_leq:
"[| l=l'; l ∈ llist(A) |] ==> <l,l'> ∈ lleq(A)"
apply (rule_tac X = "{<l,l>. l ∈ llist (A) }" in lleq.coinduct)
apply blast
apply safe
apply (erule_tac a=la in llist.cases, fast+)
done
lemma lconst_fun_bnd_mono: "bnd_mono(univ(a), %l. LCons(a,l))"
apply (unfold llist.con_defs )
apply (rule bnd_monoI)
apply (blast intro: A_subset_univ QInr_subset_univ)
apply (blast intro: subset_refl QInr_mono QPair_mono)
done
lemmas lconst = def_lfp_unfold [OF lconst_def lconst_fun_bnd_mono]
lemmas lconst_subset = lconst_def [THEN def_lfp_subset]
lemmas member_subset_Union_eclose = arg_into_eclose [THEN Union_upper]
lemma lconst_in_quniv: "a ∈ A ==> lconst(a) ∈ quniv(A)"
apply (rule lconst_subset [THEN subset_trans, THEN qunivI])
apply (erule arg_into_eclose [THEN eclose_subset, THEN univ_mono])
done
lemma lconst_type: "a ∈ A ==> lconst(a): llist(A)"
apply (rule singletonI [THEN llist.coinduct])
apply (erule lconst_in_quniv [THEN singleton_subsetI])
apply (fast intro!: lconst)
done
declare flip_LNil [simp]
flip_LCons [simp]
not_type [simp]
lemma bool_Int_subset_univ: "b ∈ bool ==> b ∩ X ⊆ univ(eclose(A))"
by (fast intro: Int_lower1 [THEN subset_trans] elim!: boolE)
declare not_type [intro!]
declare bool_Int_subset_univ [intro]
lemma flip_llist_quniv_lemma:
"Ord(i) ==> l ∈ llist(bool) ⟹ flip(l) ∩ Vset(i) ⊆ univ(eclose(bool))"
proof (induct i arbitrary: l rule: trans_induct)
case (step i l)
show ?case using ‹l ∈ llist(bool)›
proof (cases l rule: llist.cases)
case LNil thus ?thesis
by (simp, simp add: QInl_def QInr_def llist.con_defs)
next
case (LCons a l) thus ?thesis using step.hyps
proof (simp, simp add: QInl_def QInr_def llist.con_defs)
show "<1; <not(a); flip(l)>> ∩ Vset(i) ⊆ univ(eclose(bool))" using LCons step.hyps
by (auto intro: Ord_trans)
qed
qed
qed
lemma flip_in_quniv: "l ∈ llist(bool) ==> flip(l) ∈ quniv(bool)"
by (rule flip_llist_quniv_lemma [THEN Int_Vset_subset, THEN qunivI], assumption+)
lemma flip_type: "l ∈ llist(bool) ==> flip(l): llist(bool)"
apply (rule_tac X = "{flip (l) . l ∈ llist (bool) }" in llist.coinduct)
apply blast
apply (fast intro!: flip_in_quniv)
apply (erule RepFunE)
apply (erule_tac a=la in llist.cases, auto)
done
lemma flip_flip: "l ∈ llist(bool) ==> flip(flip(l)) = l"
apply (rule_tac X1 = "{<flip (flip (l)),l> . l ∈ llist (bool) }" in
lleq.coinduct [THEN lleq_implies_equal])
apply blast
apply (fast intro!: flip_type)
apply (erule RepFunE)
apply (erule_tac a=la in llist.cases)
apply (simp_all add: flip_type not_not)
done
end