--- a/Attic/Quot/Examples/FSet3.thy Wed Apr 21 09:48:35 2010 +0200
+++ b/Attic/Quot/Examples/FSet3.thy Wed Apr 21 10:13:17 2010 +0200
@@ -1,257 +1,22 @@
theory FSet3
-imports "../Quotient" "../Quotient_List" List
+imports "../../../Nominal/FSet"
begin
-ML {*
-structure QuotientRules = Named_Thms
- (val name = "quot_thm"
- val description = "Quotient theorems.")
-*}
-
-ML {*
-open QuotientRules
-*}
-
-fun
- list_eq :: "'a list \<Rightarrow> 'a list \<Rightarrow> bool" (infix "\<approx>" 50)
-where
- "list_eq xs ys = (\<forall>x. x \<in> set xs \<longleftrightarrow> x \<in> set ys)"
-
-lemma list_eq_equivp:
- shows "equivp list_eq"
-unfolding equivp_reflp_symp_transp reflp_def symp_def transp_def
-by auto
-
-(* FIXME-TODO: because of beta-reduction, one cannot give the *)
-(* FIXME-TODO: relation as a term or abbreviation *)
-quotient_type
- 'a fset = "'a list" / "list_eq"
-by (rule list_eq_equivp)
-
-
-section {* empty fset, finsert and membership *}
-
-quotient_definition
- fempty ("{||}")
-where
- "fempty :: 'a fset"
-is "[]::'a list"
-
-quotient_definition
- "finsert :: 'a \<Rightarrow> 'a fset \<Rightarrow> 'a fset"
-is "op #"
-
-syntax
- "@Finset" :: "args => 'a fset" ("{|(_)|}")
-
-translations
- "{|x, xs|}" == "CONST finsert x {|xs|}"
- "{|x|}" == "CONST finsert x {||}"
-
-definition
- memb :: "'a \<Rightarrow> 'a list \<Rightarrow> bool"
-where
- "memb x xs \<equiv> x \<in> set xs"
-
-quotient_definition
- fin ("_ |\<in>| _" [50, 51] 50)
-where
- "fin :: 'a \<Rightarrow> 'a fset \<Rightarrow> bool"
-is "memb"
-
-abbreviation
- fnotin :: "'a \<Rightarrow> 'a fset \<Rightarrow> bool" ("_ |\<notin>| _" [50, 51] 50)
-where
- "a |\<notin>| S \<equiv> \<not>(a |\<in>| S)"
-
-lemma memb_rsp[quot_respect]:
- shows "(op = ===> op \<approx> ===> op =) memb memb"
-by (auto simp add: memb_def)
-
-lemma nil_rsp[quot_respect]:
- shows "[] \<approx> []"
-by simp
-
-lemma cons_rsp:
- "xa \<approx> ya \<Longrightarrow> y # xa \<approx> y # ya"
- by simp
-
-lemma [quot_respect]:
- shows "(op = ===> op \<approx> ===> op \<approx>) op # op #"
- by simp
-
-
-section {* Augmenting a set -- @{const finsert} *}
-
-text {* raw section *}
-
-lemma nil_not_cons:
- shows "\<not>[] \<approx> x # xs"
- by auto
-
-lemma memb_cons_iff:
- shows "memb x (y # xs) = (x = y \<or> memb x xs)"
- by (induct xs) (auto simp add: memb_def)
-
-lemma memb_consI1:
- shows "memb x (x # xs)"
- by (simp add: memb_def)
-
-lemma memb_consI2:
- shows "memb x xs \<Longrightarrow> memb x (y # xs)"
- by (simp add: memb_def)
-
-lemma memb_absorb:
- shows "memb x xs \<Longrightarrow> x # xs \<approx> xs"
- by (induct xs) (auto simp add: memb_def id_simps)
-
-text {* lifted section *}
-
-lemma fin_finsert_iff[simp]:
- "x |\<in>| finsert y S = (x = y \<or> x |\<in>| S)"
-by (lifting memb_cons_iff)
-
-lemma
- shows finsertI1: "x |\<in>| finsert x S"
- and finsertI2: "x |\<in>| S \<Longrightarrow> x |\<in>| finsert y S"
- by (lifting memb_consI1, lifting memb_consI2)
-
-
-lemma finsert_absorb [simp]:
- shows "x |\<in>| S \<Longrightarrow> finsert x S = S"
- by (lifting memb_absorb)
-
-
-section {* Singletons *}
-
-text {* raw section *}
-
-lemma singleton_list_eq:
- shows "[x] \<approx> [y] \<longleftrightarrow> x = y"
- by (simp add: id_simps) auto
-
-text {* lifted section *}
-
-lemma fempty_not_finsert[simp]:
- shows "{||} \<noteq> finsert x S"
- by (lifting nil_not_cons)
-
-lemma fsingleton_eq[simp]:
- shows "{|x|} = {|y|} \<longleftrightarrow> x = y"
- by (lifting singleton_list_eq)
-
-section {* Union *}
-
-quotient_definition
- funion (infixl "|\<union>|" 65)
-where
- "funion :: 'a fset \<Rightarrow> 'a fset \<Rightarrow> 'a fset"
-is
- "op @"
-
-section {* Cardinality of finite sets *}
-
-fun
- fcard_raw :: "'a list \<Rightarrow> nat"
-where
- fcard_raw_nil: "fcard_raw [] = 0"
-| fcard_raw_cons: "fcard_raw (x # xs) = (if memb x xs then fcard_raw xs else Suc (fcard_raw xs))"
-
-quotient_definition
- "fcard :: 'a fset \<Rightarrow> nat"
-is "fcard_raw"
-
-text {* raw section *}
-
-lemma fcard_raw_ge_0:
- assumes a: "x \<in> set xs"
- shows "0 < fcard_raw xs"
-using a
-by (induct xs) (auto simp add: memb_def)
-
-lemma fcard_raw_delete_one:
- "fcard_raw ([x \<leftarrow> xs. x \<noteq> y]) = (if memb y xs then fcard_raw xs - 1 else fcard_raw xs)"
-by (induct xs) (auto dest: fcard_raw_ge_0 simp add: memb_def)
-
-lemma fcard_raw_rsp_aux:
- assumes a: "a \<approx> b"
- shows "fcard_raw a = fcard_raw b"
-using a
-apply(induct a arbitrary: b)
-apply(auto simp add: memb_def)
-apply(metis)
-apply(drule_tac x="[x \<leftarrow> b. x \<noteq> a1]" in meta_spec)
-apply(simp add: fcard_raw_delete_one)
-apply(metis Suc_pred' fcard_raw_ge_0 fcard_raw_delete_one memb_def)
-done
-
-lemma fcard_raw_rsp[quot_respect]:
- "(op \<approx> ===> op =) fcard_raw fcard_raw"
- by (simp add: fcard_raw_rsp_aux)
-
-text {* lifted section *}
-
-lemma fcard_fempty [simp]:
- shows "fcard {||} = 0"
-by (lifting fcard_raw_nil)
-
-lemma fcard_finsert_if [simp]:
- shows "fcard (finsert x S) = (if x |\<in>| S then fcard S else Suc (fcard S))"
-by (lifting fcard_raw_cons)
-
-
-section {* Induction and Cases rules for finite sets *}
+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 fset_induct[case_names fempty finsert]:
- shows "\<lbrakk>P {||}; \<And>x S. P S \<Longrightarrow> P (finsert x S)\<rbrakk> \<Longrightarrow> P S"
-by (lifting list.induct)
-
-lemma fset_induct2[case_names fempty finsert, induct type: fset]:
- assumes prem1: "P {||}"
- and prem2: "\<And>x S. \<lbrakk>x |\<notin>| S; P S\<rbrakk> \<Longrightarrow> P (finsert x S)"
- shows "P S"
-proof(induct S rule: fset_induct)
- case fempty
- show "P {||}" by (rule prem1)
-next
- case (finsert x S)
- have asm: "P S" by fact
- show "P (finsert x S)"
- proof(cases "x |\<in>| S")
- case True
- have "x |\<in>| S" by fact
- then show "P (finsert x S)" using asm by simp
- next
- case False
- have "x |\<notin>| S" by fact
- then show "P (finsert x S)" using prem2 asm by simp
- qed
-qed
-
-
-section {* fmap and fset comprehension *}
-
-quotient_definition
- "fmap :: ('a \<Rightarrow> 'b) \<Rightarrow> 'a fset \<Rightarrow> 'b fset"
-is
- "map"
-
-quotient_definition
- "fconcat :: ('a fset) fset => 'a fset"
-is
- "concat"
-
(* PROBLEM: these lemmas needs to be restated, since *)
(* concat.simps(1) and concat.simps(2) contain the *)
(* type variables ?'a1.0 (which are turned into frees *)
(* 'a_1 *)
+
lemma concat1:
shows "concat [] \<approx> []"
-by (simp add: id_simps)
+by (simp)
lemma concat2:
shows "concat (x # xs) \<approx> x @ concat xs"
@@ -274,8 +39,7 @@
done
lemma nil_rsp2[quot_respect]: "(list_rel op \<approx> OOO op \<approx>) [] []"
- apply (metis FSet3.nil_rsp list_rel.simps(1) pred_comp.intros)
- done
+ 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"
apply (rule eq_reflection)
@@ -361,9 +125,7 @@
shows "fconcat {||} = {||}"
apply(lifting concat1)
apply(cleaning)
- apply(simp add: comp_def)
- apply(fold fempty_def[simplified id_simps])
- apply(rule refl)
+ apply(simp add: comp_def bot_fset_def)
done
lemma insert_rsp2[quot_respect]:
@@ -373,9 +135,7 @@
apply (rule_tac b="x # b" in pred_compI)
apply auto
apply (rule_tac b="x # ba" in pred_compI)
- apply (rule cons_rsp)
- apply simp
- apply (auto)[1]
+ apply auto
done
lemma append_rsp[quot_respect]: