theory FSet3
imports "../../../Nominal/FSet"
begin
notation
list_eq (infix "\<approx>" 50)
lemma fset_exhaust[case_names fempty finsert, cases type: fset]:
shows "\<lbrakk>S = {||} \<Longrightarrow> P; \<And>x S'. S = finsert x S' \<Longrightarrow> P\<rbrakk> \<Longrightarrow> P"
by (lifting list.exhaust)
lemma list_rel_find_element:
assumes a: "x \<in> set a"
and b: "list_rel R a b"
shows "\<exists>y. (y \<in> set b \<and> R x y)"
proof -
have "length a = length b" using b by (rule list_rel_len)
then show ?thesis using a b proof (induct a b rule: list_induct2)
case Nil
show ?case using Nil.prems by simp
next
case (Cons x xs y ys)
show ?case using Cons by auto
qed
qed
lemma concat_rsp_pre:
assumes a: "list_rel op \<approx> x x'"
and b: "x' \<approx> y'"
and c: "list_rel op \<approx> y' y"
and d: "\<exists>x\<in>set x. xa \<in> set x"
shows "\<exists>x\<in>set y. xa \<in> set x"
proof -
obtain xb where e: "xb \<in> set x" and f: "xa \<in> set xb" using d by auto
have "\<exists>y. y \<in> set x' \<and> xb \<approx> y" by (rule list_rel_find_element[OF e a])
then obtain ya where h: "ya \<in> set x'" and i: "xb \<approx> ya" by auto
have j: "ya \<in> set y'" using b h by simp
have "\<exists>yb. yb \<in> set y \<and> ya \<approx> yb" by (rule list_rel_find_element[OF j c])
then show ?thesis using f i by auto
qed
lemma fun_relI [intro]:
assumes "\<And>a b. P a b \<Longrightarrow> Q (x a) (y b)"
shows "(P ===> Q) x y"
using assms by (simp add: fun_rel_def)
lemma [quot_respect]:
shows "(list_rel op \<approx> OOO op \<approx> ===> op \<approx>) concat concat"
proof (rule fun_relI, elim pred_compE)
fix a b ba bb
assume a: "list_rel op \<approx> a ba"
assume b: "ba \<approx> bb"
assume c: "list_rel op \<approx> bb b"
have "\<forall>x. (\<exists>xa\<in>set a. x \<in> set xa) = (\<exists>xa\<in>set b. x \<in> set xa)" proof
fix x
show "(\<exists>xa\<in>set a. x \<in> set xa) = (\<exists>xa\<in>set b. x \<in> set xa)" proof
assume d: "\<exists>xa\<in>set a. x \<in> set xa"
show "\<exists>xa\<in>set b. x \<in> set xa" by (rule concat_rsp_pre[OF a b c d])
next
assume e: "\<exists>xa\<in>set b. x \<in> set xa"
have a': "list_rel op \<approx> ba a" by (rule list_rel_symp[OF list_eq_equivp, OF a])
have b': "bb \<approx> ba" by (rule equivp_symp[OF list_eq_equivp, OF b])
have c': "list_rel op \<approx> b bb" by (rule list_rel_symp[OF list_eq_equivp, OF c])
show "\<exists>xa\<in>set a. x \<in> set xa" by (rule concat_rsp_pre[OF c' b' a' e])
qed
qed
then show "concat a \<approx> concat b" by simp
qed
lemma nil_rsp2[quot_respect]: "(list_rel op \<approx> OOO op \<approx>) [] []"
by (metis nil_rsp list_rel.simps(1) pred_compI)
lemma set_in_eq: "(\<forall>e. ((e \<in> A) \<longleftrightarrow> (e \<in> B))) \<equiv> A = B"
by (rule eq_reflection) auto
lemma map_rel_cong: "b \<approx> ba \<Longrightarrow> map f b \<approx> map f ba"
unfolding list_eq.simps
by (simp only: set_map set_in_eq)
lemma compose_list_refl:
shows "(list_rel op \<approx> OOO op \<approx>) r r"
proof
show c: "list_rel op \<approx> r r" by (rule list_rel_refl) (metis equivp_def fset_equivp)
have d: "r \<approx> r" by (rule equivp_reflp[OF fset_equivp])
show b: "(op \<approx> OO list_rel op \<approx>) r r" by (rule pred_compI) (rule d, rule c)
qed
lemma list_rel_refl:
shows "(list_rel op \<approx>) r r"
by (rule list_rel_refl)(metis equivp_def fset_equivp)
lemma Quotient_fset_list:
shows "Quotient (list_rel op \<approx>) (map abs_fset) (map rep_fset)"
by (fact list_quotient[OF Quotient_fset])
lemma quotient_compose_list[quot_thm]:
shows "Quotient ((list_rel op \<approx>) OOO (op \<approx>))
(abs_fset \<circ> (map abs_fset)) ((map rep_fset) \<circ> rep_fset)"
unfolding Quotient_def comp_def
proof (intro conjI allI)
fix a r s
show "abs_fset (map abs_fset (map rep_fset (rep_fset a))) = a"
by (simp add: abs_o_rep[OF Quotient_fset] Quotient_abs_rep[OF Quotient_fset] map_id)
have b: "list_rel op \<approx> (map rep_fset (rep_fset a)) (map rep_fset (rep_fset a))"
by (rule list_rel_refl)
have c: "(op \<approx> OO list_rel op \<approx>) (map rep_fset (rep_fset a)) (map rep_fset (rep_fset a))"
by (rule, rule equivp_reflp[OF fset_equivp]) (rule b)
show "(list_rel op \<approx> OOO op \<approx>) (map rep_fset (rep_fset a)) (map rep_fset (rep_fset a))"
by (rule, rule list_rel_refl) (rule c)
show "(list_rel op \<approx> OOO op \<approx>) r s = ((list_rel op \<approx> OOO op \<approx>) r r \<and>
(list_rel op \<approx> OOO op \<approx>) s s \<and> abs_fset (map abs_fset r) = abs_fset (map abs_fset s))"
proof (intro iffI conjI)
show "(list_rel op \<approx> OOO op \<approx>) r r" by (rule compose_list_refl)
show "(list_rel op \<approx> OOO op \<approx>) s s" by (rule compose_list_refl)
next
assume a: "(list_rel op \<approx> OOO op \<approx>) r s"
then have b: "map abs_fset r \<approx> map abs_fset s" proof (elim pred_compE)
fix b ba
assume c: "list_rel op \<approx> r b"
assume d: "b \<approx> ba"
assume e: "list_rel op \<approx> ba s"
have f: "map abs_fset r = map abs_fset b"
by (metis Quotient_rel[OF Quotient_fset_list] c)
have g: "map abs_fset s = map abs_fset ba"
by (metis Quotient_rel[OF Quotient_fset_list] e)
show "map abs_fset r \<approx> map abs_fset s" using d f g map_rel_cong by simp
qed
then show "abs_fset (map abs_fset r) = abs_fset (map abs_fset s)"
by (metis Quotient_rel[OF Quotient_fset])
next
assume a: "(list_rel op \<approx> OOO op \<approx>) r r \<and> (list_rel op \<approx> OOO op \<approx>) s s
\<and> abs_fset (map abs_fset r) = abs_fset (map abs_fset s)"
then have s: "(list_rel op \<approx> OOO op \<approx>) s s" by simp
have d: "map abs_fset r \<approx> map abs_fset s"
by (subst Quotient_rel[OF Quotient_fset]) (simp add: a)
have b: "map rep_fset (map abs_fset r) \<approx> map rep_fset (map abs_fset s)"
by (rule map_rel_cong[OF d])
have y: "list_rel op \<approx> (map rep_fset (map abs_fset s)) s"
by (fact rep_abs_rsp_left[OF Quotient_fset_list, OF list_rel_refl[of s]])
have c: "(op \<approx> OO list_rel op \<approx>) (map rep_fset (map abs_fset r)) s"
by (rule pred_compI) (rule b, rule y)
have z: "list_rel op \<approx> r (map rep_fset (map abs_fset r))"
by (fact rep_abs_rsp[OF Quotient_fset_list, OF list_rel_refl[of r]])
then show "(list_rel op \<approx> OOO op \<approx>) r s"
using a c pred_compI by simp
qed
qed
lemma nil_prs2[quot_preserve]: "(abs_fset \<circ> map f) [] = abs_fset []"
by simp
lemma fconcat_empty:
shows "fconcat {||} = {||}"
by (lifting concat.simps(1))
lemma insert_rsp2[quot_respect]:
"(op \<approx> ===> list_rel op \<approx> OOO op \<approx> ===> list_rel op \<approx> OOO op \<approx>) op # op #"
apply auto
apply (simp add: set_in_eq)
apply (rule_tac b="x # b" in pred_compI)
apply auto
apply (rule_tac b="x # ba" in pred_compI)
apply auto
done
lemma append_rsp[quot_respect]:
"(op \<approx> ===> op \<approx> ===> op \<approx>) op @ op @"
by (auto)
lemma insert_prs2[quot_preserve]:
"(rep_fset ---> (map rep_fset \<circ> rep_fset) ---> (abs_fset \<circ> map abs_fset)) op # = finsert"
by (simp add: expand_fun_eq Quotient_abs_rep[OF Quotient_fset]
abs_o_rep[OF Quotient_fset] map_id finsert_def)
lemma fconcat_insert:
shows "fconcat (finsert x S) = x |\<union>| fconcat S"
by (lifting concat.simps(2))
lemma append_prs2[quot_preserve]:
"((map rep_fset \<circ> rep_fset) ---> (map rep_fset \<circ> rep_fset) ---> abs_fset \<circ> map abs_fset) op @ = funion"
by (simp add: expand_fun_eq Quotient_abs_rep[OF Quotient_fset]
abs_o_rep[OF Quotient_fset] map_id sup_fset_def)
lemma list_rel_app_l:
assumes a: "reflp R"
and b: "list_rel R l r"
shows "list_rel R (z @ l) (z @ r)"
by (induct z) (simp_all add: b, metis a reflp_def)
lemma append_rsp2_pre0:
assumes a:"list_rel op \<approx> x x'"
shows "list_rel op \<approx> (x @ z) (x' @ z)"
using a apply (induct x x' rule: list_induct2')
apply simp_all
apply (rule list_rel_refl)
done
lemma append_rsp2_pre1:
assumes a:"list_rel op \<approx> x x'"
shows "list_rel op \<approx> (z @ x) (z @ x')"
using a apply (induct x x' arbitrary: z rule: list_induct2')
apply (rule list_rel_refl)
apply (simp_all del: list_eq.simps)
apply (rule list_rel_app_l)
apply (simp_all add: reflp_def)
done
lemma append_rsp2_pre:
assumes a:"list_rel op \<approx> x x'"
and b: "list_rel op \<approx> z z'"
shows "list_rel op \<approx> (x @ z) (x' @ z')"
apply (rule list_rel_transp[OF fset_equivp])
apply (rule append_rsp2_pre0)
apply (rule a)
using b apply (induct z z' rule: list_induct2')
apply (simp_all only: append_Nil2)
apply (rule list_rel_refl)
apply simp_all
apply (rule append_rsp2_pre1)
apply simp
done
lemma append_rsp2[quot_respect]:
"(list_rel op \<approx> OOO op \<approx> ===> list_rel op \<approx> OOO op \<approx> ===> list_rel op \<approx> OOO op \<approx>) op @ op @"
proof (intro fun_relI, elim pred_compE)
fix x y z w x' z' y' w' :: "'a list list"
assume a:"list_rel op \<approx> x x'"
and b: "x' \<approx> y'"
and c: "list_rel op \<approx> y' y"
assume aa: "list_rel op \<approx> z z'"
and bb: "z' \<approx> w'"
and cc: "list_rel op \<approx> w' w"
have a': "list_rel op \<approx> (x @ z) (x' @ z')" using a aa append_rsp2_pre by auto
have b': "x' @ z' \<approx> y' @ w'" using b bb by simp
have c': "list_rel op \<approx> (y' @ w') (y @ w)" using c cc append_rsp2_pre by auto
have d': "(op \<approx> OO list_rel op \<approx>) (x' @ z') (y @ w)"
by (rule pred_compI) (rule b', rule c')
show "(list_rel op \<approx> OOO op \<approx>) (x @ z) (y @ w)"
by (rule pred_compI) (rule a', rule d')
qed
lemma "fconcat (xs |\<union>| ys) = fconcat xs |\<union>| fconcat ys"
by (lifting concat_append)
(* TBD *)
text {* syntax for fset comprehensions (adapted from lists) *}
nonterminals fsc_qual fsc_quals
syntax
"_fsetcompr" :: "'a \<Rightarrow> fsc_qual \<Rightarrow> fsc_quals \<Rightarrow> 'a fset" ("{|_ . __")
"_fsc_gen" :: "'a \<Rightarrow> 'a fset \<Rightarrow> fsc_qual" ("_ <- _")
"_fsc_test" :: "bool \<Rightarrow> fsc_qual" ("_")
"_fsc_end" :: "fsc_quals" ("|}")
"_fsc_quals" :: "fsc_qual \<Rightarrow> fsc_quals \<Rightarrow> fsc_quals" (", __")
"_fsc_abs" :: "'a => 'b fset => 'b fset"
syntax (xsymbols)
"_fsc_gen" :: "'a \<Rightarrow> 'a fset \<Rightarrow> fsc_qual" ("_ \<leftarrow> _")
syntax (HTML output)
"_fsc_gen" :: "'a \<Rightarrow> 'a fset \<Rightarrow> fsc_qual" ("_ \<leftarrow> _")
ML {* Syntax.check_term @{context} (Const ("List.append", dummyT) $ @{term "[3::nat,4]"}) *}
parse_translation (advanced) {*
let
val femptyC = Syntax.const @{const_name fempty};
val finsertC = Syntax.const @{const_name finsert};
val fmapC = Syntax.const @{const_name fmap};
val fconcatC = Syntax.const @{const_name fconcat};
val IfC = Syntax.const @{const_name If};
fun fsingl x = finsertC $ x $ femptyC;
fun pat_tr ctxt p e opti = (* %x. case x of p => e | _ => [] *)
let
val x = Free (Name.variant (fold Term.add_free_names [p, e] []) "x", dummyT);
val e = if opti then fsingl e else e;
val case1 = Syntax.const "_case1" $ p $ e;
val case2 = Syntax.const "_case1" $ Syntax.const Term.dummy_patternN
$ femptyC;
val cs = Syntax.const "_case2" $ case1 $ case2
val ft = Datatype_Case.case_tr false Datatype.info_of_constr
ctxt [x, cs]
in lambda x ft end;
fun abs_tr ctxt (p as Free(s,T)) e opti =
let val thy = ProofContext.theory_of ctxt;
val s' = Sign.intern_const thy s
in if Sign.declared_const thy s'
then (pat_tr ctxt p e opti, false)
else (lambda p e, true)
end
| abs_tr ctxt p e opti = (pat_tr ctxt p e opti, false);
fun fsc_tr ctxt [e, Const("_fsc_test",_) $ b, qs] =
let
val res = case qs of
Const("_fsc_end",_) => fsingl e
| Const("_fsc_quals",_)$ q $ qs => fsc_tr ctxt [e, q, qs];
in
IfC $ b $ res $ femptyC
end
| fsc_tr ctxt [e, Const("_fsc_gen",_) $ p $ es, Const("_fsc_end",_)] =
(case abs_tr ctxt p e true of
(f,true) => fmapC $ f $ es
| (f, false) => fconcatC $ (fmapC $ f $ es))
| fsc_tr ctxt [e, Const("_fsc_gen",_) $ p $ es, Const("_fsc_quals",_) $ q $ qs] =
let
val e' = fsc_tr ctxt [e, q, qs];
in
fconcatC $ (fmapC $ (fst (abs_tr ctxt p e' false)) $ es)
end
in [("_fsetcompr", fsc_tr)] end
*}
(* NEEDS FIXING *)
(* examles *)
(*
term "{|(x,y,z). b|}"
term "{|x. x \<leftarrow> xs|}"
term "{|(x,y,z). x\<leftarrow>xs|}"
term "{|e x y. x\<leftarrow>xs, y\<leftarrow>ys|}"
term "{|(x,y,z). x<a, x>b|}"
term "{|(x,y,z). x\<leftarrow>xs, x>b|}"
term "{|(x,y,z). x<a, x\<leftarrow>xs|}"
term "{|(x,y). Cons True x \<leftarrow> xs|}"
term "{|(x,y,z). Cons x [] \<leftarrow> xs|}"
term "{|(x,y,z). x<a, x>b, x=d|}"
term "{|(x,y,z). x<a, x>b, y\<leftarrow>ys|}"
term "{|(x,y,z). x<a, x\<leftarrow>xs,y>b|}"
term "{|(x,y,z). x<a, x\<leftarrow>xs, y\<leftarrow>ys|}"
term "{|(x,y,z). x\<leftarrow>xs, x>b, y<a|}"
term "{|(x,y,z). x\<leftarrow>xs, x>b, y\<leftarrow>ys|}"
term "{|(x,y,z). x\<leftarrow>xs, y\<leftarrow>ys,y>x|}"
term "{|(x,y,z). x\<leftarrow>xs, y\<leftarrow>ys,z\<leftarrow>zs|}"
*)
(* BELOW CONSTRUCTION SITE *)
end