FSet.thy
changeset 597 8a1c8dc72b5c
parent 596 6088fea1c8b1
child 598 ae254a6d685c
equal deleted inserted replaced
596:6088fea1c8b1 597:8a1c8dc72b5c
     1 theory FSet
       
     2 imports QuotMain
       
     3 begin
       
     4 
       
     5 inductive
       
     6   list_eq (infix "\<approx>" 50)
       
     7 where
       
     8   "a#b#xs \<approx> b#a#xs"
       
     9 | "[] \<approx> []"
       
    10 | "xs \<approx> ys \<Longrightarrow> ys \<approx> xs"
       
    11 | "a#a#xs \<approx> a#xs"
       
    12 | "xs \<approx> ys \<Longrightarrow> a#xs \<approx> a#ys"
       
    13 | "\<lbrakk>xs1 \<approx> xs2; xs2 \<approx> xs3\<rbrakk> \<Longrightarrow> xs1 \<approx> xs3"
       
    14 
       
    15 lemma list_eq_refl:
       
    16   shows "xs \<approx> xs"
       
    17   by (induct xs) (auto intro: list_eq.intros)
       
    18 
       
    19 lemma equivp_list_eq:
       
    20   shows "equivp list_eq"
       
    21   unfolding equivp_reflp_symp_transp reflp_def symp_def transp_def
       
    22   apply(auto intro: list_eq.intros list_eq_refl)
       
    23   done
       
    24 
       
    25 quotient fset = "'a list" / "list_eq"
       
    26   apply(rule equivp_list_eq)
       
    27   done
       
    28 
       
    29 print_theorems
       
    30 
       
    31 typ "'a fset"
       
    32 thm "Rep_fset"
       
    33 thm "ABS_fset_def"
       
    34 
       
    35 quotient_def 
       
    36   EMPTY :: "'a fset"
       
    37 where
       
    38   "EMPTY \<equiv> ([]::'a list)"
       
    39 
       
    40 term Nil
       
    41 term EMPTY
       
    42 thm EMPTY_def
       
    43 
       
    44 quotient_def 
       
    45   INSERT :: "'a \<Rightarrow> 'a fset \<Rightarrow> 'a fset"
       
    46 where
       
    47   "INSERT \<equiv> op #"
       
    48 
       
    49 term Cons
       
    50 term INSERT
       
    51 thm INSERT_def
       
    52 
       
    53 quotient_def 
       
    54   FUNION :: "'a fset \<Rightarrow> 'a fset \<Rightarrow> 'a fset"
       
    55 where
       
    56   "FUNION \<equiv> (op @)"
       
    57 
       
    58 term append
       
    59 term FUNION
       
    60 thm FUNION_def
       
    61 
       
    62 thm Quotient_fset
       
    63 
       
    64 thm QUOT_TYPE_I_fset.thm11
       
    65 
       
    66 
       
    67 fun
       
    68   membship :: "'a \<Rightarrow> 'a list \<Rightarrow> bool" (infix "memb" 100)
       
    69 where
       
    70   m1: "(x memb []) = False"
       
    71 | m2: "(x memb (y#xs)) = ((x=y) \<or> (x memb xs))"
       
    72 
       
    73 fun
       
    74   card1 :: "'a list \<Rightarrow> nat"
       
    75 where
       
    76   card1_nil: "(card1 []) = 0"
       
    77 | card1_cons: "(card1 (x # xs)) = (if (x memb xs) then (card1 xs) else (Suc (card1 xs)))"
       
    78 
       
    79 quotient_def 
       
    80   CARD :: "'a fset \<Rightarrow> nat"
       
    81 where
       
    82   "CARD \<equiv> card1"
       
    83 
       
    84 term card1
       
    85 term CARD
       
    86 thm CARD_def
       
    87 
       
    88 (* text {*
       
    89  Maybe make_const_def should require a theorem that says that the particular lifted function
       
    90  respects the relation. With it such a definition would be impossible:
       
    91  make_const_def @{binding CARD} @{term "length"} NoSyn @{typ "'a list"} @{typ "'a fset"} #> snd
       
    92 *}*)
       
    93 
       
    94 lemma card1_0:
       
    95   fixes a :: "'a list"
       
    96   shows "(card1 a = 0) = (a = [])"
       
    97   by (induct a) auto
       
    98 
       
    99 lemma not_mem_card1:
       
   100   fixes x :: "'a"
       
   101   fixes xs :: "'a list"
       
   102   shows "(~(x memb xs)) = (card1 (x # xs) = Suc (card1 xs))"
       
   103   by auto
       
   104 
       
   105 lemma mem_cons:
       
   106   fixes x :: "'a"
       
   107   fixes xs :: "'a list"
       
   108   assumes a : "x memb xs"
       
   109   shows "x # xs \<approx> xs"
       
   110   using a by (induct xs) (auto intro: list_eq.intros )
       
   111 
       
   112 lemma card1_suc:
       
   113   fixes xs :: "'a list"
       
   114   fixes n :: "nat"
       
   115   assumes c: "card1 xs = Suc n"
       
   116   shows "\<exists>a ys. ~(a memb ys) \<and> xs \<approx> (a # ys)"
       
   117   using c
       
   118 apply(induct xs)
       
   119 apply (metis Suc_neq_Zero card1_0)
       
   120 apply (metis QUOT_TYPE_I_fset.R_trans card1_cons list_eq_refl mem_cons)
       
   121 done
       
   122 
       
   123 definition
       
   124   rsp_fold
       
   125 where
       
   126   "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))))"
       
   127 
       
   128 primrec
       
   129   fold1
       
   130 where
       
   131   "fold1 f (g :: 'a \<Rightarrow> 'b) (z :: 'b) [] = z"
       
   132 | "fold1 f g z (a # A) =
       
   133      (if rsp_fold f
       
   134      then (
       
   135        if (a memb A) then (fold1 f g z A) else (f (g a) (fold1 f g z A))
       
   136      ) else z)"
       
   137 
       
   138 lemma fs1_strong_cases:
       
   139   fixes X :: "'a list"
       
   140   shows "(X = []) \<or> (\<exists>a. \<exists> Y. (~(a memb Y) \<and> (X \<approx> a # Y)))"
       
   141   apply (induct X)
       
   142   apply (simp)
       
   143   apply (metis QUOT_TYPE_I_fset.thm11 list_eq_refl mem_cons m1)
       
   144   done
       
   145 
       
   146 quotient_def
       
   147   IN :: "'a \<Rightarrow> 'a fset \<Rightarrow> bool"
       
   148 where
       
   149   "IN \<equiv> membship"
       
   150 
       
   151 term membship
       
   152 term IN
       
   153 thm IN_def
       
   154 
       
   155 term fold1
       
   156 quotient_def 
       
   157   FOLD :: "('a \<Rightarrow> 'a \<Rightarrow> 'a) \<Rightarrow> ('b \<Rightarrow> 'a) \<Rightarrow> 'a \<Rightarrow> 'b fset \<Rightarrow> 'a"
       
   158 where
       
   159   "FOLD \<equiv> fold1"
       
   160 
       
   161 term fold1
       
   162 term fold
       
   163 thm fold_def
       
   164 
       
   165 quotient_def 
       
   166   fmap::"('a \<Rightarrow> 'b) \<Rightarrow> 'a fset \<Rightarrow> 'b fset"
       
   167 where
       
   168   "fmap \<equiv> map"
       
   169 
       
   170 term map
       
   171 term fmap
       
   172 thm fmap_def
       
   173 
       
   174 lemma memb_rsp:
       
   175   fixes z
       
   176   assumes a: "x \<approx> y"
       
   177   shows "(z memb x) = (z memb y)"
       
   178   using a by induct auto
       
   179 
       
   180 lemma ho_memb_rsp[quotient_rsp]:
       
   181   "(op = ===> (op \<approx> ===> op =)) (op memb) (op memb)"
       
   182   by (simp add: memb_rsp)
       
   183 
       
   184 lemma card1_rsp:
       
   185   fixes a b :: "'a list"
       
   186   assumes e: "a \<approx> b"
       
   187   shows "card1 a = card1 b"
       
   188   using e by induct (simp_all add:memb_rsp)
       
   189 
       
   190 lemma ho_card1_rsp[quotient_rsp]: 
       
   191   "(op \<approx> ===> op =) card1 card1"
       
   192   by (simp add: card1_rsp)
       
   193 
       
   194 lemma cons_rsp[quotient_rsp]:
       
   195   fixes z
       
   196   assumes a: "xs \<approx> ys"
       
   197   shows "(z # xs) \<approx> (z # ys)"
       
   198   using a by (rule list_eq.intros(5))
       
   199 
       
   200 lemma ho_cons_rsp[quotient_rsp]:
       
   201   "(op = ===> op \<approx> ===> op \<approx>) op # op #"
       
   202   by (simp add: cons_rsp)
       
   203 
       
   204 lemma append_rsp_fst:
       
   205   assumes a : "l1 \<approx> l2"
       
   206   shows "(l1 @ s) \<approx> (l2 @ s)"
       
   207   using a
       
   208   by (induct) (auto intro: list_eq.intros list_eq_refl)
       
   209 
       
   210 lemma append_end:
       
   211   shows "(e # l) \<approx> (l @ [e])"
       
   212   apply (induct l)
       
   213   apply (auto intro: list_eq.intros list_eq_refl)
       
   214   done
       
   215 
       
   216 lemma rev_rsp:
       
   217   shows "a \<approx> rev a"
       
   218   apply (induct a)
       
   219   apply simp
       
   220   apply (rule list_eq_refl)
       
   221   apply simp_all
       
   222   apply (rule list_eq.intros(6))
       
   223   prefer 2
       
   224   apply (rule append_rsp_fst)
       
   225   apply assumption
       
   226   apply (rule append_end)
       
   227   done
       
   228 
       
   229 lemma append_sym_rsp:
       
   230   shows "(a @ b) \<approx> (b @ a)"
       
   231   apply (rule list_eq.intros(6))
       
   232   apply (rule append_rsp_fst)
       
   233   apply (rule rev_rsp)
       
   234   apply (rule list_eq.intros(6))
       
   235   apply (rule rev_rsp)
       
   236   apply (simp)
       
   237   apply (rule append_rsp_fst)
       
   238   apply (rule list_eq.intros(3))
       
   239   apply (rule rev_rsp)
       
   240   done
       
   241 
       
   242 lemma append_rsp:
       
   243   assumes a : "l1 \<approx> r1"
       
   244   assumes b : "l2 \<approx> r2 "
       
   245   shows "(l1 @ l2) \<approx> (r1 @ r2)"
       
   246   apply (rule list_eq.intros(6))
       
   247   apply (rule append_rsp_fst)
       
   248   using a apply (assumption)
       
   249   apply (rule list_eq.intros(6))
       
   250   apply (rule append_sym_rsp)
       
   251   apply (rule list_eq.intros(6))
       
   252   apply (rule append_rsp_fst)
       
   253   using b apply (assumption)
       
   254   apply (rule append_sym_rsp)
       
   255   done
       
   256 
       
   257 lemma ho_append_rsp[quotient_rsp]:
       
   258   "(op \<approx> ===> op \<approx> ===> op \<approx>) op @ op @"
       
   259   by (simp add: append_rsp)
       
   260 
       
   261 lemma map_rsp:
       
   262   assumes a: "a \<approx> b"
       
   263   shows "map f a \<approx> map f b"
       
   264   using a
       
   265   apply (induct)
       
   266   apply(auto intro: list_eq.intros)
       
   267   done
       
   268 
       
   269 lemma ho_map_rsp[quotient_rsp]:
       
   270   "(op = ===> op \<approx> ===> op \<approx>) map map"
       
   271   by (simp add: map_rsp)
       
   272 
       
   273 lemma map_append:
       
   274   "(map f (a @ b)) \<approx> (map f a) @ (map f b)"
       
   275  by simp (rule list_eq_refl)
       
   276 
       
   277 lemma ho_fold_rsp[quotient_rsp]:
       
   278   "(op = ===> op = ===> op = ===> op \<approx> ===> op =) fold1 fold1"
       
   279   apply (auto)
       
   280   apply (case_tac "rsp_fold x")
       
   281   prefer 2
       
   282   apply (erule_tac list_eq.induct)
       
   283   apply (simp_all)
       
   284   apply (erule_tac list_eq.induct)
       
   285   apply (simp_all)
       
   286   apply (auto simp add: memb_rsp rsp_fold_def)
       
   287 done
       
   288 
       
   289 lemma list_equiv_rsp[quotient_rsp]:
       
   290   shows "(op \<approx> ===> op \<approx> ===> op =) op \<approx> op \<approx>"
       
   291 by (auto intro: list_eq.intros)
       
   292 
       
   293 print_quotients
       
   294 
       
   295 ML {* val qty = @{typ "'a fset"} *}
       
   296 ML {* val rsp_thms =
       
   297   @{thms ho_memb_rsp ho_cons_rsp ho_card1_rsp ho_map_rsp ho_append_rsp ho_fold_rsp} *}
       
   298 
       
   299 ML {* val (rty, rel, rel_refl, rel_eqv) = lookup_quot_data @{context} qty *}
       
   300 ML {* val (trans2, reps_same, absrep, quot) = lookup_quot_thms @{context} "fset"; *}
       
   301 ML {* fun lift_tac_fset lthy t = lift_tac lthy t *}
       
   302 
       
   303 lemma "IN x EMPTY = False"
       
   304 apply(tactic {* procedure_tac @{context} @{thm m1} 1 *})
       
   305 apply(tactic {* regularize_tac @{context} 1 *})
       
   306 apply(tactic {* all_inj_repabs_tac @{context} [rel_refl] [trans2] 1 *})
       
   307 apply(tactic {* clean_tac @{context} 1*})
       
   308 done
       
   309 
       
   310 lemma "IN x (INSERT y xa) = (x = y \<or> IN x xa)"
       
   311 by (tactic {* lift_tac_fset @{context} @{thm m2} 1 *})
       
   312 
       
   313 lemma "INSERT a (INSERT a x) = INSERT a x"
       
   314 apply (tactic {* lift_tac_fset @{context} @{thm list_eq.intros(4)} 1 *})
       
   315 done
       
   316 
       
   317 lemma "x = xa \<Longrightarrow> INSERT a x = INSERT a xa"
       
   318 apply (tactic {* lift_tac_fset @{context} @{thm list_eq.intros(5)} 1 *})
       
   319 done
       
   320 
       
   321 lemma "CARD x = Suc n \<Longrightarrow> (\<exists>a b. \<not> IN a b & x = INSERT a b)"
       
   322 apply (tactic {* lift_tac_fset @{context} @{thm card1_suc} 1 *})
       
   323 done
       
   324 
       
   325 lemma "(\<not> IN x xa) = (CARD (INSERT x xa) = Suc (CARD xa))"
       
   326 apply (tactic {* lift_tac_fset @{context} @{thm not_mem_card1} 1 *})
       
   327 done
       
   328 
       
   329 lemma "FOLD f g (z::'b) (INSERT a x) =
       
   330   (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)"
       
   331 apply(tactic {* lift_tac_fset @{context} @{thm fold1.simps(2)} 1 *})
       
   332 done
       
   333 
       
   334 ML {* fun inj_repabs_tac_fset lthy = inj_repabs_tac lthy [rel_refl] [trans2] *}
       
   335 
       
   336 lemma "fmap f (FUNION (x::'b fset) (xa::'b fset)) = FUNION (fmap f x) (fmap f xa)"
       
   337 apply (tactic {* lift_tac_fset @{context} @{thm map_append} 1 *})
       
   338 done
       
   339 
       
   340 lemma "FUNION (FUNION x xa) xb = FUNION x (FUNION xa xb)"
       
   341 apply (tactic {* lift_tac_fset @{context} @{thm append_assoc} 1 *})
       
   342 done
       
   343 
       
   344 
       
   345 lemma "\<lbrakk>P EMPTY; \<And>a x. P x \<Longrightarrow> P (INSERT a x)\<rbrakk> \<Longrightarrow> P l"
       
   346 apply (tactic {* (ObjectLogic.full_atomize_tac THEN' gen_frees_tac @{context}) 1 *})
       
   347 apply(tactic {* procedure_tac @{context} @{thm list.induct} 1 *})
       
   348 apply(tactic {* regularize_tac @{context} 1 *})
       
   349 defer
       
   350 apply(tactic {* clean_tac @{context} 1 *})
       
   351 apply(tactic {* inj_repabs_tac_fset @{context} 1*})+
       
   352 done
       
   353 
       
   354 lemma list_induct_part:
       
   355   assumes a: "P (x :: 'a list) ([] :: 'c list)"
       
   356   assumes b: "\<And>e t. P x t \<Longrightarrow> P x (e # t)"
       
   357   shows "P x l"
       
   358   apply (rule_tac P="P x" in list.induct)
       
   359   apply (rule a)
       
   360   apply (rule b)
       
   361   apply (assumption)
       
   362   done
       
   363 
       
   364 ML {* quot *}
       
   365 thm quotient_thm
       
   366 
       
   367 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"
       
   368 apply (tactic {* lift_tac_fset @{context} @{thm list_induct_part} 1 *})
       
   369 done
       
   370 
       
   371 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"
       
   372 apply (tactic {* lift_tac_fset @{context} @{thm list_induct_part} 1 *})
       
   373 done
       
   374 
       
   375 lemma "P (x :: 'a fset) ([] :: 'c list) \<Longrightarrow> (\<And>e t. P x t \<Longrightarrow> P x (e # t)) \<Longrightarrow> P x l"
       
   376 apply (tactic {* lift_tac_fset @{context} @{thm list_induct_part} 1 *})
       
   377 done
       
   378 
       
   379 quotient fset2 = "'a list" / "list_eq"
       
   380   apply(rule equivp_list_eq)
       
   381   done
       
   382 
       
   383 quotient_def
       
   384   EMPTY2 :: "'a fset2"
       
   385 where
       
   386   "EMPTY2 \<equiv> ([]::'a list)"
       
   387 
       
   388 quotient_def
       
   389   INSERT2 :: "'a \<Rightarrow> 'a fset2 \<Rightarrow> 'a fset2"
       
   390 where
       
   391   "INSERT2 \<equiv> op #"
       
   392 
       
   393 ML {* val quot = @{thms Quotient_fset Quotient_fset2} *}
       
   394 ML {* fun inj_repabs_tac_fset lthy = inj_repabs_tac lthy [rel_refl] [trans2] *}
       
   395 ML {* fun lift_tac_fset lthy t = lift_tac lthy t  *}
       
   396 
       
   397 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"
       
   398 apply (tactic {* lift_tac_fset @{context} @{thm list_induct_part} 1 *})
       
   399 done
       
   400 
       
   401 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"
       
   402 apply (tactic {* lift_tac_fset @{context} @{thm list_induct_part} 1 *})
       
   403 done
       
   404 
       
   405 quotient_def
       
   406   fset_rec::"'a \<Rightarrow> ('b \<Rightarrow> 'b fset \<Rightarrow> 'a \<Rightarrow> 'a) \<Rightarrow> 'b fset \<Rightarrow> 'a"
       
   407 where
       
   408   "fset_rec \<equiv> list_rec"
       
   409 
       
   410 quotient_def
       
   411   fset_case::"'a \<Rightarrow> ('b \<Rightarrow> 'b fset \<Rightarrow> 'a) \<Rightarrow> 'b fset \<Rightarrow> 'a"
       
   412 where
       
   413   "fset_case \<equiv> list_case"
       
   414 
       
   415 (* Probably not true without additional assumptions about the function *)
       
   416 lemma list_rec_rsp[quotient_rsp]:
       
   417   "(op = ===> (op = ===> op \<approx> ===> op =) ===> op \<approx> ===> op =) list_rec list_rec"
       
   418   apply (auto)
       
   419   apply (erule_tac list_eq.induct)
       
   420   apply (simp_all)
       
   421   sorry
       
   422 
       
   423 lemma list_case_rsp[quotient_rsp]:
       
   424   "(op = ===> (op = ===> op \<approx> ===> op =) ===> op \<approx> ===> op =) list_case list_case"
       
   425   apply (auto)
       
   426   sorry
       
   427 
       
   428 ML {* val rsp_thms = @{thms list_rec_rsp list_case_rsp} @ rsp_thms *}
       
   429 ML {* fun lift_tac_fset lthy t = lift_tac lthy t *}
       
   430 
       
   431 lemma "fset_rec (f1::'t) x (INSERT a xa) = x a xa (fset_rec f1 x xa)"
       
   432 apply (tactic {* lift_tac_fset @{context} @{thm list.recs(2)} 1 *})
       
   433 done
       
   434 
       
   435 lemma "fset_case (f1::'t) f2 (INSERT a xa) = f2 a xa"
       
   436 apply (tactic {* lift_tac_fset @{context} @{thm list.cases(2)} 1 *})
       
   437 done
       
   438 
       
   439 
       
   440 end