theory FSet+ −
imports "../QuotMain" List+ −
begin+ −
+ −
inductive+ −
list_eq (infix "\<approx>" 50)+ −
where+ −
"a#b#xs \<approx> b#a#xs"+ −
| "[] \<approx> []"+ −
| "xs \<approx> ys \<Longrightarrow> ys \<approx> xs"+ −
| "a#a#xs \<approx> a#xs"+ −
| "xs \<approx> ys \<Longrightarrow> a#xs \<approx> a#ys"+ −
| "\<lbrakk>xs1 \<approx> xs2; xs2 \<approx> xs3\<rbrakk> \<Longrightarrow> xs1 \<approx> xs3"+ −
+ −
lemma list_eq_refl:+ −
shows "xs \<approx> xs"+ −
by (induct xs) (auto intro: list_eq.intros)+ −
+ −
lemma equivp_list_eq:+ −
shows "equivp list_eq"+ −
unfolding equivp_reflp_symp_transp reflp_def symp_def transp_def+ −
apply(auto intro: list_eq.intros list_eq_refl)+ −
done+ −
+ −
quotient fset = "'a list" / "list_eq"+ −
apply(rule equivp_list_eq)+ −
done+ −
+ −
quotient_def+ −
EMPTY :: "EMPTY :: 'a fset"+ −
where+ −
"[]::'a list"+ −
+ −
quotient_def+ −
INSERT :: "INSERT :: 'a \<Rightarrow> 'a fset \<Rightarrow> 'a fset"+ −
where+ −
"op #"+ −
+ −
quotient_def+ −
FUNION :: "FUNION :: 'a fset \<Rightarrow> 'a fset \<Rightarrow> 'a fset"+ −
where+ −
"op @"+ −
+ −
fun+ −
membship :: "'a \<Rightarrow> 'a list \<Rightarrow> bool" (infix "memb" 100)+ −
where+ −
m1: "(x memb []) = False"+ −
| m2: "(x memb (y#xs)) = ((x=y) \<or> (x memb xs))"+ −
+ −
fun+ −
card1 :: "'a list \<Rightarrow> nat"+ −
where+ −
card1_nil: "(card1 []) = 0"+ −
| card1_cons: "(card1 (x # xs)) = (if (x memb xs) then (card1 xs) else (Suc (card1 xs)))"+ −
+ −
quotient_def+ −
CARD :: "CARD :: 'a fset \<Rightarrow> nat"+ −
where+ −
"card1"+ −
+ −
(* text {*+ −
Maybe make_const_def should require a theorem that says that the particular lifted function+ −
respects the relation. With it such a definition would be impossible:+ −
make_const_def @{binding CARD} @{term "length"} NoSyn @{typ "'a list"} @{typ "'a fset"} #> snd+ −
*}*)+ −
+ −
lemma card1_0:+ −
fixes a :: "'a list"+ −
shows "(card1 a = 0) = (a = [])"+ −
by (induct a) auto+ −
+ −
lemma not_mem_card1:+ −
fixes x :: "'a"+ −
fixes xs :: "'a list"+ −
shows "(~(x memb xs)) = (card1 (x # xs) = Suc (card1 xs))"+ −
by auto+ −
+ −
lemma mem_cons:+ −
fixes x :: "'a"+ −
fixes xs :: "'a list"+ −
assumes a : "x memb xs"+ −
shows "x # xs \<approx> xs"+ −
using a by (induct xs) (auto intro: list_eq.intros )+ −
+ −
lemma card1_suc:+ −
fixes xs :: "'a list"+ −
fixes n :: "nat"+ −
assumes c: "card1 xs = Suc n"+ −
shows "\<exists>a ys. ~(a memb ys) \<and> xs \<approx> (a # ys)"+ −
using c+ −
apply(induct xs)+ −
apply (metis Suc_neq_Zero card1_0)+ −
apply (metis QUOT_TYPE_I_fset.R_trans card1_cons list_eq_refl mem_cons)+ −
done+ −
+ −
definition+ −
rsp_fold+ −
where+ −
"rsp_fold f = ((!u v. (f u v = f v u)) \<and> (!u v w. ((f u (f v w) = f (f u v) w))))"+ −
+ −
primrec+ −
fold1+ −
where+ −
"fold1 f (g :: 'a \<Rightarrow> 'b) (z :: 'b) [] = z"+ −
| "fold1 f g z (a # A) =+ −
(if rsp_fold f+ −
then (+ −
if (a memb A) then (fold1 f g z A) else (f (g a) (fold1 f g z A))+ −
) else z)"+ −
+ −
lemma fs1_strong_cases:+ −
fixes X :: "'a list"+ −
shows "(X = []) \<or> (\<exists>a. \<exists> Y. (~(a memb Y) \<and> (X \<approx> a # Y)))"+ −
apply (induct X)+ −
apply (simp)+ −
apply (metis QUOT_TYPE_I_fset.thm11 list_eq_refl mem_cons m1)+ −
done+ −
+ −
quotient_def+ −
IN :: "IN :: 'a \<Rightarrow> 'a fset \<Rightarrow> bool"+ −
where+ −
"membship"+ −
+ −
quotient_def+ −
FOLD :: "FOLD :: ('a \<Rightarrow> 'a \<Rightarrow> 'a) \<Rightarrow> ('b \<Rightarrow> 'a) \<Rightarrow> 'a \<Rightarrow> 'b fset \<Rightarrow> 'a"+ −
where+ −
"fold1"+ −
+ −
quotient_def+ −
fmap::"fmap :: ('a \<Rightarrow> 'b) \<Rightarrow> 'a fset \<Rightarrow> 'b fset"+ −
where+ −
"map"+ −
+ −
lemma memb_rsp:+ −
fixes z+ −
assumes a: "x \<approx> y"+ −
shows "(z memb x) = (z memb y)"+ −
using a by induct auto+ −
+ −
lemma ho_memb_rsp[quot_respect]:+ −
"(op = ===> (op \<approx> ===> op =)) (op memb) (op memb)"+ −
by (simp add: memb_rsp)+ −
+ −
lemma card1_rsp:+ −
fixes a b :: "'a list"+ −
assumes e: "a \<approx> b"+ −
shows "card1 a = card1 b"+ −
using e by induct (simp_all add:memb_rsp)+ −
+ −
lemma ho_card1_rsp[quot_respect]: + −
"(op \<approx> ===> op =) card1 card1"+ −
by (simp add: card1_rsp)+ −
+ −
lemma cons_rsp:+ −
fixes z+ −
assumes a: "xs \<approx> ys"+ −
shows "(z # xs) \<approx> (z # ys)"+ −
using a by (rule list_eq.intros(5))+ −
+ −
lemma ho_cons_rsp[quot_respect]:+ −
"(op = ===> op \<approx> ===> op \<approx>) op # op #"+ −
by (simp add: cons_rsp)+ −
+ −
lemma append_rsp_aux1:+ −
assumes a : "l2 \<approx> r2 "+ −
shows "(h @ l2) \<approx> (h @ r2)"+ −
using a+ −
apply(induct h)+ −
apply(auto intro: list_eq.intros(5))+ −
done+ −
+ −
lemma append_rsp_aux2:+ −
assumes a : "l1 \<approx> r1" "l2 \<approx> r2 "+ −
shows "(l1 @ l2) \<approx> (r1 @ r2)"+ −
using a+ −
apply(induct arbitrary: l2 r2)+ −
apply(simp_all)+ −
apply(blast intro: list_eq.intros append_rsp_aux1)++ −
done+ −
+ −
lemma append_rsp[quot_respect]:+ −
"(op \<approx> ===> op \<approx> ===> op \<approx>) op @ op @"+ −
by (auto simp add: append_rsp_aux2)+ −
+ −
lemma map_rsp:+ −
assumes a: "a \<approx> b"+ −
shows "map f a \<approx> map f b"+ −
using a+ −
apply (induct)+ −
apply(auto intro: list_eq.intros)+ −
done+ −
+ −
lemma ho_map_rsp[quot_respect]:+ −
"(op = ===> op \<approx> ===> op \<approx>) map map"+ −
by (simp add: map_rsp)+ −
+ −
lemma map_append:+ −
"(map f (a @ b)) \<approx> (map f a) @ (map f b)"+ −
by simp (rule list_eq_refl)+ −
+ −
lemma ho_fold_rsp[quot_respect]:+ −
"(op = ===> op = ===> op = ===> op \<approx> ===> op =) fold1 fold1"+ −
apply (auto)+ −
apply (case_tac "rsp_fold x")+ −
prefer 2+ −
apply (erule_tac list_eq.induct)+ −
apply (simp_all)+ −
apply (erule_tac list_eq.induct)+ −
apply (simp_all)+ −
apply (auto simp add: memb_rsp rsp_fold_def)+ −
done+ −
+ −
lemma list_equiv_rsp[quot_respect]:+ −
shows "(op \<approx> ===> op \<approx> ===> op =) op \<approx> op \<approx>"+ −
by (auto intro: list_eq.intros)+ −
+ −
+ −
lemma "IN x EMPTY = False"+ −
apply(lifting m1)+ −
done+ −
+ −
lemma "IN x (INSERT y xa) = (x = y \<or> IN x xa)"+ −
by (lifting m2)+ −
+ −
lemma "INSERT a (INSERT a x) = INSERT a x"+ −
apply (lifting list_eq.intros(4))+ −
done+ −
+ −
lemma "x = xa \<Longrightarrow> INSERT a x = INSERT a xa"+ −
apply (lifting list_eq.intros(5))+ −
done+ −
+ −
lemma "CARD x = Suc n \<Longrightarrow> (\<exists>a b. \<not> IN a b & x = INSERT a b)"+ −
apply (lifting card1_suc)+ −
done+ −
+ −
lemma "(\<not> IN x xa) = (CARD (INSERT x xa) = Suc (CARD xa))"+ −
apply (lifting not_mem_card1)+ −
done+ −
+ −
lemma "FOLD f g (z::'b) (INSERT a x) =+ −
(if rsp_fold f then if IN a x then FOLD f g z x else f (g a) (FOLD f g z x) else z)"+ −
apply(lifting fold1.simps(2))+ −
done+ −
+ −
lemma "fmap f (FUNION (x::'b fset) (xa::'b fset)) = FUNION (fmap f x) (fmap f xa)"+ −
apply (lifting map_append)+ −
done+ −
+ −
lemma "FUNION (FUNION x xa) xb = FUNION x (FUNION xa xb)"+ −
apply (lifting append_assoc)+ −
done+ −
+ −
+ −
lemma "\<lbrakk>P EMPTY; \<And>a x. P x \<Longrightarrow> P (INSERT a x)\<rbrakk> \<Longrightarrow> P l"+ −
apply(lifting list.induct)+ −
done+ −
+ −
lemma list_induct_part:+ −
assumes a: "P (x :: 'a list) ([] :: 'c list)"+ −
assumes b: "\<And>e t. P x t \<Longrightarrow> P x (e # t)"+ −
shows "P x l"+ −
apply (rule_tac P="P x" in list.induct)+ −
apply (rule a)+ −
apply (rule b)+ −
apply (assumption)+ −
done+ −
+ −
lemma "P (x :: 'a list) (EMPTY :: 'c fset) \<Longrightarrow> (\<And>e t. P x t \<Longrightarrow> P x (INSERT e t)) \<Longrightarrow> P x l"+ −
apply (lifting list_induct_part)+ −
done+ −
+ −
lemma "P (x :: 'a fset) (EMPTY :: 'c fset) \<Longrightarrow> (\<And>e t. P x t \<Longrightarrow> P x (INSERT e t)) \<Longrightarrow> P x l"+ −
apply (lifting list_induct_part)+ −
done+ −
+ −
lemma "P (x :: 'a fset) ([] :: 'c list) \<Longrightarrow> (\<And>e t. P x t \<Longrightarrow> P x (e # t)) \<Longrightarrow> P x l"+ −
apply (lifting list_induct_part)+ −
done+ −
+ −
quotient fset2 = "'a list" / "list_eq"+ −
by (rule equivp_list_eq)+ −
+ −
quotient_def+ −
EMPTY2 :: "EMPTY2 :: 'a fset2"+ −
where+ −
"[]::'a list"+ −
+ −
quotient_def+ −
INSERT2 :: "INSERT2 :: 'a \<Rightarrow> 'a fset2 \<Rightarrow> 'a fset2"+ −
where+ −
"op #"+ −
+ −
lemma "P (x :: 'a fset2) (EMPTY :: 'c fset) \<Longrightarrow> (\<And>e t. P x t \<Longrightarrow> P x (INSERT e t)) \<Longrightarrow> P x l"+ −
apply (lifting list_induct_part)+ −
done+ −
+ −
lemma "P (x :: 'a fset) (EMPTY2 :: 'c fset2) \<Longrightarrow> (\<And>e t. P x t \<Longrightarrow> P x (INSERT2 e t)) \<Longrightarrow> P x l"+ −
apply (lifting list_induct_part)+ −
done+ −
+ −
quotient_def+ −
fset_rec::"fset_rec :: 'a \<Rightarrow> ('b \<Rightarrow> 'b fset \<Rightarrow> 'a \<Rightarrow> 'a) \<Rightarrow> 'b fset \<Rightarrow> 'a"+ −
where+ −
"list_rec"+ −
+ −
quotient_def+ −
fset_case::"fset_case :: 'a \<Rightarrow> ('b \<Rightarrow> 'b fset \<Rightarrow> 'a) \<Rightarrow> 'b fset \<Rightarrow> 'a"+ −
where+ −
"list_case"+ −
+ −
(* Probably not true without additional assumptions about the function *)+ −
lemma list_rec_rsp[quot_respect]:+ −
"(op = ===> (op = ===> op \<approx> ===> op =) ===> op \<approx> ===> op =) list_rec list_rec"+ −
apply (auto)+ −
apply (erule_tac list_eq.induct)+ −
apply (simp_all)+ −
sorry+ −
+ −
lemma list_case_rsp[quot_respect]:+ −
"(op = ===> (op = ===> op \<approx> ===> op =) ===> op \<approx> ===> op =) list_case list_case"+ −
apply (auto)+ −
sorry+ −
+ −
lemma "fset_rec (f1::'t) x (INSERT a xa) = x a xa (fset_rec f1 x xa)"+ −
apply (lifting list.recs(2))+ −
done+ −
+ −
lemma "fset_case (f1::'t) f2 (INSERT a xa) = f2 a xa"+ −
apply (lifting list.cases(2))+ −
done+ −
+ −
lemma ttt: "((op @) x ((op #) e [])) = (((op #) e x))"+ −
sorry+ −
+ −
lemma "(FUNION x (INSERT e EMPTY)) = ((INSERT e x))"+ −
apply (lifting ttt)+ −
done+ −
+ −
+ −
lemma ttt2: "(\<lambda>e. ((op @) x ((op #) e []))) = (\<lambda>e. ((op #) e x))"+ −
sorry+ −
+ −
lemma "(\<lambda>e. (FUNION x (INSERT e EMPTY))) = (\<lambda>e. (INSERT e x))"+ −
apply(lifting ttt2)+ −
apply(regularize)+ −
apply(rule impI)+ −
apply(simp)+ −
apply(rule allI)+ −
apply(rule list_eq_refl)+ −
done+ −
+ −
lemma ttt3: "(\<lambda>x. ((op @) x ((op #) e []))) = (\<lambda>x. ((op #) e x))"+ −
sorry+ −
+ −
+ −
lemma "(\<lambda>x. (FUNION x (INSERT e EMPTY))) = (\<lambda>x. (INSERT e x))"+ −
(* apply (tactic {* procedure_tac @{context} @{thm ttt3} 1 *}) *)+ −
sorry+ −
+ −
lemma hard: "(\<lambda>P. \<lambda>Q. P (Q (x::'a list))) = (\<lambda>P. \<lambda>Q. Q (P (x::'a list)))"+ −
sorry+ −
+ −
lemma hard_lift: "(\<lambda>P. \<lambda>Q. P (Q (x::'a fset))) = (\<lambda>P. \<lambda>Q. Q (P (x::'a fset)))"+ −
apply(lifting hard)+ −
apply(regularize)+ −
apply(auto)+ −
sorry+ −
+ −
end+ −