| 0 |      1 | theory QuotMain
 | 
|  |      2 | imports QuotScript QuotList
 | 
|  |      3 | begin
 | 
|  |      4 | 
 | 
|  |      5 | locale QUOT_TYPE =
 | 
|  |      6 |   fixes R :: "'a \<Rightarrow> 'a \<Rightarrow> bool"
 | 
|  |      7 |   and   Abs :: "('a \<Rightarrow> bool) \<Rightarrow> 'b"
 | 
|  |      8 |   and   Rep :: "'b \<Rightarrow> ('a \<Rightarrow> bool)"
 | 
|  |      9 |   assumes equiv: "EQUIV R"
 | 
|  |     10 |   and     rep_prop: "\<And>y. \<exists>x. Rep y = R x"
 | 
|  |     11 |   and     rep_inverse: "\<And>x. Abs (Rep x) = x"
 | 
|  |     12 |   and     abs_inverse: "\<And>x. (Rep (Abs (R x))) = (R x)"
 | 
|  |     13 |   and     rep_inject: "\<And>x y. (Rep x = Rep y) = (x = y)"
 | 
|  |     14 | begin 
 | 
|  |     15 | 
 | 
|  |     16 | definition
 | 
|  |     17 |   "ABS x \<equiv> Abs (R x)"
 | 
|  |     18 | 
 | 
|  |     19 | definition
 | 
|  |     20 |   "REP a = Eps (Rep a)"
 | 
|  |     21 | 
 | 
|  |     22 | lemma lem9: 
 | 
|  |     23 |   shows "R (Eps (R x)) = R x"
 | 
|  |     24 | proof -
 | 
|  |     25 |   have a: "R x x" using equiv by (simp add: EQUIV_REFL_SYM_TRANS REFL_def)
 | 
|  |     26 |   then have "R x (Eps (R x))" by (rule someI)
 | 
|  |     27 |   then show "R (Eps (R x)) = R x" 
 | 
|  |     28 |     using equiv unfolding EQUIV_def by simp
 | 
|  |     29 | qed
 | 
|  |     30 | 
 | 
|  |     31 | theorem thm10:
 | 
|  |     32 |   shows "ABS (REP a) = a"
 | 
|  |     33 | unfolding ABS_def REP_def
 | 
|  |     34 | proof -
 | 
|  |     35 |   from rep_prop 
 | 
|  |     36 |   obtain x where eq: "Rep a = R x" by auto
 | 
|  |     37 |   have "Abs (R (Eps (Rep a))) = Abs (R (Eps (R x)))" using eq by simp
 | 
|  |     38 |   also have "\<dots> = Abs (R x)" using lem9 by simp
 | 
|  |     39 |   also have "\<dots> = Abs (Rep a)" using eq by simp
 | 
|  |     40 |   also have "\<dots> = a" using rep_inverse by simp
 | 
|  |     41 |   finally
 | 
|  |     42 |   show "Abs (R (Eps (Rep a))) = a" by simp
 | 
|  |     43 | qed
 | 
|  |     44 | 
 | 
|  |     45 | lemma REP_refl: 
 | 
|  |     46 |   shows "R (REP a) (REP a)"
 | 
|  |     47 | unfolding REP_def
 | 
|  |     48 | by (simp add: equiv[simplified EQUIV_def])
 | 
|  |     49 | 
 | 
|  |     50 | lemma lem7:
 | 
|  |     51 |   "(R x = R y) = (Abs (R x) = Abs (R y))"
 | 
|  |     52 | apply(rule iffI)
 | 
|  |     53 | apply(simp)
 | 
|  |     54 | apply(drule rep_inject[THEN iffD2])
 | 
|  |     55 | apply(simp add: abs_inverse)
 | 
|  |     56 | done
 | 
|  |     57 |   
 | 
|  |     58 | theorem thm11:
 | 
|  |     59 |   shows "R r r' = (ABS r = ABS r')"
 | 
|  |     60 | unfolding ABS_def
 | 
|  |     61 | by (simp only: equiv[simplified EQUIV_def] lem7)
 | 
|  |     62 | 
 | 
| 2 |     63 | lemma REP_ABS_rsp:
 | 
|  |     64 |   shows "R f g \<Longrightarrow> R f (REP (ABS g))"
 | 
|  |     65 | apply(subst thm11)
 | 
|  |     66 | apply(simp add: thm10 thm11)
 | 
|  |     67 | done
 | 
|  |     68 | 
 | 
| 0 |     69 | lemma QUOTIENT:
 | 
|  |     70 |   "QUOTIENT R ABS REP"
 | 
|  |     71 | apply(unfold QUOTIENT_def)
 | 
|  |     72 | apply(simp add: thm10)
 | 
|  |     73 | apply(simp add: REP_refl)
 | 
|  |     74 | apply(subst thm11[symmetric])
 | 
|  |     75 | apply(simp add: equiv[simplified EQUIV_def])
 | 
|  |     76 | done
 | 
|  |     77 | 
 | 
|  |     78 | end
 | 
|  |     79 | 
 | 
|  |     80 | section {* type definition for the quotient type *}
 | 
|  |     81 | 
 | 
|  |     82 | ML {*
 | 
| 2 |     83 | Variable.variant_frees
 | 
|  |     84 | *}
 | 
|  |     85 | 
 | 
|  |     86 | 
 | 
|  |     87 | ML {*
 | 
| 1 |     88 | (* constructs the term \<lambda>(c::ty \<Rightarrow> bool). \<exists>x. c = rel x *)
 | 
| 0 |     89 | fun typedef_term rel ty lthy =
 | 
|  |     90 | let 
 | 
|  |     91 |   val [x, c] = [("x", ty), ("c", ty --> @{typ bool})] 
 | 
|  |     92 |                |> Variable.variant_frees lthy [rel]
 | 
|  |     93 |                |> map Free
 | 
|  |     94 | in
 | 
| 1 |     95 |   lambda c 
 | 
|  |     96 |     (HOLogic.mk_exists 
 | 
|  |     97 |        ("x", ty, HOLogic.mk_eq (c, (rel $ x))))
 | 
| 0 |     98 | end
 | 
|  |     99 | *}
 | 
|  |    100 | 
 | 
|  |    101 | ML {*
 | 
| 2 |    102 |  typedef_term @{term R} @{typ "nat"} @{context} 
 | 
|  |    103 |   |> Syntax.string_of_term @{context}
 | 
|  |    104 |   |> writeln
 | 
|  |    105 | *}
 | 
|  |    106 | 
 | 
|  |    107 | ML {*
 | 
| 0 |    108 | val typedef_tac =  
 | 
|  |    109 |   EVERY1 [rewrite_goal_tac @{thms mem_def},
 | 
|  |    110 |           rtac @{thm exI}, rtac @{thm exI}, rtac @{thm refl}]
 | 
|  |    111 | *}
 | 
|  |    112 | 
 | 
|  |    113 | ML {*
 | 
|  |    114 | (* makes the new type definitions *)
 | 
|  |    115 | fun typedef_make (qty_name, rel, ty) lthy =
 | 
|  |    116 |   LocalTheory.theory_result
 | 
| 1 |    117 |   (Typedef.add_typedef false NONE 
 | 
| 0 |    118 |      (qty_name, map fst (Term.add_tfreesT ty []), NoSyn)
 | 
|  |    119 |        (typedef_term rel ty lthy)
 | 
|  |    120 |          NONE typedef_tac) lthy
 | 
|  |    121 | *}
 | 
|  |    122 | 
 | 
| 2 |    123 | text {* proves the QUOT_TYPE theorem for the new type *}
 | 
| 0 |    124 | ML {*
 | 
| 1 |    125 | fun typedef_quot_type_tac equiv_thm (typedef_info: Typedef.info) = 
 | 
| 0 |    126 | let
 | 
|  |    127 |   val rep_thm = #Rep typedef_info
 | 
|  |    128 |   val rep_inv = #Rep_inverse typedef_info
 | 
|  |    129 |   val abs_inv = #Abs_inverse typedef_info
 | 
|  |    130 |   val rep_inj = #Rep_inject typedef_info
 | 
|  |    131 | 
 | 
|  |    132 |   val ss = HOL_basic_ss addsimps @{thms mem_def}
 | 
|  |    133 |   val rep_thm_simpd = Simplifier.asm_full_simplify ss rep_thm
 | 
|  |    134 |   val abs_inv_simpd = Simplifier.asm_full_simplify ss abs_inv
 | 
|  |    135 | in
 | 
|  |    136 |   EVERY1 [rtac @{thm QUOT_TYPE.intro},
 | 
|  |    137 |           rtac equiv_thm, 
 | 
|  |    138 |           rtac rep_thm_simpd, 
 | 
|  |    139 |           rtac rep_inv, 
 | 
|  |    140 |           rtac abs_inv_simpd, rtac @{thm exI}, rtac @{thm refl},
 | 
|  |    141 |           rtac rep_inj]
 | 
|  |    142 | end
 | 
|  |    143 | *}
 | 
|  |    144 | 
 | 
| 2 |    145 | term QUOT_TYPE
 | 
|  |    146 | ML {* HOLogic.mk_Trueprop *}
 | 
|  |    147 | ML {* Goal.prove *}
 | 
|  |    148 | ML {* Syntax.check_term *}
 | 
|  |    149 | 
 | 
| 0 |    150 | ML {* 
 | 
|  |    151 | fun typedef_quot_type_thm (rel, abs, rep, equiv_thm, typedef_info) lthy =
 | 
|  |    152 | let
 | 
|  |    153 |   val quot_type_const = Const (@{const_name "QUOT_TYPE"}, dummyT)
 | 
|  |    154 |   val goal = HOLogic.mk_Trueprop (quot_type_const $ rel $ abs $ rep)
 | 
|  |    155 |              |> Syntax.check_term lthy
 | 
|  |    156 | in
 | 
|  |    157 |   Goal.prove lthy [] [] goal
 | 
|  |    158 |     (fn _ => typedef_quot_type_tac equiv_thm typedef_info)
 | 
|  |    159 | end
 | 
|  |    160 | *}
 | 
|  |    161 | 
 | 
|  |    162 | ML {*
 | 
|  |    163 | fun typedef_quotient_thm_tac defs quot_type_thm =
 | 
|  |    164 |   EVERY1 [K (rewrite_goals_tac defs), 
 | 
|  |    165 |           rtac @{thm QUOT_TYPE.QUOTIENT}, 
 | 
|  |    166 |           rtac quot_type_thm]
 | 
|  |    167 | *}
 | 
|  |    168 | 
 | 
|  |    169 | ML {* 
 | 
|  |    170 | fun typedef_quotient_thm (rel, abs, rep, abs_def, rep_def, quot_type_thm) lthy =
 | 
|  |    171 | let
 | 
|  |    172 |   val quotient_const = Const (@{const_name "QUOTIENT"}, dummyT)
 | 
|  |    173 |   val goal = HOLogic.mk_Trueprop (quotient_const $ rel $ abs $ rep)
 | 
|  |    174 |              |> Syntax.check_term lthy
 | 
|  |    175 | in
 | 
|  |    176 |   Goal.prove lthy [] [] goal
 | 
|  |    177 |     (fn _ => typedef_quotient_thm_tac [abs_def, rep_def] quot_type_thm)
 | 
|  |    178 | end
 | 
|  |    179 | *}
 | 
|  |    180 | 
 | 
|  |    181 | text {* two wrappers for define and note *}
 | 
|  |    182 | ML {* 
 | 
|  |    183 | fun make_def (name, mx, trm) lthy =
 | 
|  |    184 | let   
 | 
|  |    185 |   val ((trm, (_ , thm)), lthy') = 
 | 
|  |    186 |      LocalTheory.define Thm.internalK ((name, mx), (Attrib.empty_binding, trm)) lthy
 | 
|  |    187 | in
 | 
|  |    188 |   ((trm, thm), lthy')
 | 
|  |    189 | end
 | 
|  |    190 | *}
 | 
|  |    191 | 
 | 
|  |    192 | ML {*
 | 
|  |    193 | fun reg_thm (name, thm) lthy = 
 | 
|  |    194 | let
 | 
|  |    195 |   val ((_,[thm']), lthy') = LocalTheory.note Thm.theoremK ((name, []), [thm]) lthy 
 | 
|  |    196 | in
 | 
|  |    197 |   (thm',lthy')
 | 
|  |    198 | end
 | 
|  |    199 | *}
 | 
|  |    200 | 
 | 
|  |    201 | ML {* 
 | 
|  |    202 | fun typedef_main (qty_name, rel, ty, equiv_thm) lthy =
 | 
|  |    203 | let 
 | 
|  |    204 |   (* generates typedef *)
 | 
|  |    205 |   val ((_,typedef_info), lthy') = typedef_make (qty_name, rel, ty) lthy
 | 
|  |    206 | 
 | 
|  |    207 |   (* abs and rep functions *)
 | 
|  |    208 |   val abs_ty = #abs_type typedef_info
 | 
|  |    209 |   val rep_ty = #rep_type typedef_info
 | 
|  |    210 |   val abs_name = #Abs_name typedef_info
 | 
|  |    211 |   val rep_name = #Rep_name typedef_info
 | 
|  |    212 |   val abs = Const (abs_name, rep_ty --> abs_ty)
 | 
|  |    213 |   val rep = Const (rep_name, abs_ty --> rep_ty)
 | 
|  |    214 | 
 | 
|  |    215 |   (* ABS and REP definitions *)
 | 
|  |    216 |   val ABS_const = Const (@{const_name "QUOT_TYPE.ABS"}, dummyT )
 | 
|  |    217 |   val REP_const = Const (@{const_name "QUOT_TYPE.REP"}, dummyT )
 | 
|  |    218 |   val ABS_trm = Syntax.check_term lthy' (ABS_const $ rel $ abs) 
 | 
|  |    219 |   val REP_trm = Syntax.check_term lthy' (REP_const $ rep) 
 | 
|  |    220 |   val ABS_name = Binding.prefix_name "ABS_" qty_name
 | 
|  |    221 |   val REP_name = Binding.prefix_name "REP_" qty_name  
 | 
|  |    222 |   val (((ABS, ABS_def), (REP, REP_def)), lthy'') = 
 | 
|  |    223 |          lthy' |> make_def (ABS_name, NoSyn, ABS_trm)
 | 
|  |    224 |                ||>> make_def (REP_name, NoSyn, REP_trm)
 | 
|  |    225 | 
 | 
|  |    226 |   (* quot_type theorem *)
 | 
|  |    227 |   val quot_thm = typedef_quot_type_thm (rel, abs, rep, equiv_thm, typedef_info) lthy''
 | 
|  |    228 |   val quot_thm_name = Binding.prefix_name "QUOT_TYPE_" qty_name
 | 
|  |    229 | 
 | 
|  |    230 |   (* quotient theorem *)
 | 
|  |    231 |   val quotient_thm = typedef_quotient_thm (rel, ABS, REP, ABS_def, REP_def, quot_thm) lthy''
 | 
|  |    232 |   val quotient_thm_name = Binding.prefix_name "QUOTIENT_" qty_name
 | 
|  |    233 | 
 | 
|  |    234 | in
 | 
|  |    235 |   lthy'' 
 | 
|  |    236 |   |> reg_thm (quot_thm_name, quot_thm)  
 | 
|  |    237 |   ||>> reg_thm (quotient_thm_name, quotient_thm)
 | 
|  |    238 | end
 | 
|  |    239 | *}
 | 
|  |    240 | 
 | 
|  |    241 | section {* various tests for quotient types*}
 | 
|  |    242 | datatype trm =
 | 
|  |    243 |   var  "nat" 
 | 
|  |    244 | | app  "trm" "trm"
 | 
|  |    245 | | lam  "nat" "trm"
 | 
|  |    246 |   
 | 
|  |    247 | consts R :: "trm \<Rightarrow> trm \<Rightarrow> bool"
 | 
|  |    248 | axioms r_eq: "EQUIV R"
 | 
|  |    249 | 
 | 
| 2 |    250 | ML {*
 | 
|  |    251 |   typedef_main 
 | 
|  |    252 | *}
 | 
|  |    253 | 
 | 
| 0 |    254 | local_setup {*
 | 
|  |    255 |   typedef_main (@{binding "qtrm"}, @{term "R"}, @{typ trm}, @{thm r_eq}) #> snd
 | 
|  |    256 | *}
 | 
|  |    257 | 
 | 
|  |    258 | term Rep_qtrm
 | 
| 2 |    259 | term REP_qtrm
 | 
| 0 |    260 | term Abs_qtrm
 | 
|  |    261 | term ABS_qtrm
 | 
|  |    262 | thm QUOT_TYPE_qtrm
 | 
|  |    263 | thm QUOTIENT_qtrm
 | 
| 2 |    264 | 
 | 
| 0 |    265 | thm Rep_qtrm
 | 
|  |    266 | 
 | 
|  |    267 | text {* another test *}
 | 
|  |    268 | datatype 'a my = foo
 | 
|  |    269 | consts Rmy :: "'a my \<Rightarrow> 'a my \<Rightarrow> bool"
 | 
|  |    270 | axioms rmy_eq: "EQUIV Rmy"
 | 
|  |    271 | 
 | 
|  |    272 | term "\<lambda>(c::'a my\<Rightarrow>bool). \<exists>x. c = Rmy x"
 | 
|  |    273 | 
 | 
|  |    274 | datatype 'a trm' =
 | 
|  |    275 |   var'  "'a" 
 | 
|  |    276 | | app'  "'a trm'" "'a trm'"
 | 
|  |    277 | | lam'  "'a" "'a trm'"
 | 
|  |    278 |   
 | 
|  |    279 | consts R' :: "'a trm' \<Rightarrow> 'a trm' \<Rightarrow> bool"
 | 
|  |    280 | axioms r_eq': "EQUIV R'"
 | 
|  |    281 | 
 | 
|  |    282 | local_setup {*
 | 
|  |    283 |   typedef_main (@{binding "qtrm'"}, @{term "R'"}, @{typ "'a trm'"}, @{thm r_eq'}) #> snd
 | 
|  |    284 | *}
 | 
|  |    285 | 
 | 
|  |    286 | term ABS_qtrm'
 | 
|  |    287 | term REP_qtrm'
 | 
|  |    288 | thm QUOT_TYPE_qtrm'
 | 
|  |    289 | thm QUOTIENT_qtrm'
 | 
|  |    290 | thm Rep_qtrm'
 | 
|  |    291 | 
 | 
|  |    292 | text {* a test with lists of terms *}
 | 
|  |    293 | datatype t =
 | 
|  |    294 |   vr "string"
 | 
|  |    295 | | ap "t list"
 | 
|  |    296 | | lm "string" "t"
 | 
|  |    297 | 
 | 
|  |    298 | consts Rt :: "t \<Rightarrow> t \<Rightarrow> bool"
 | 
|  |    299 | axioms t_eq: "EQUIV Rt"
 | 
|  |    300 | 
 | 
|  |    301 | local_setup {*
 | 
|  |    302 |   typedef_main (@{binding "qt"}, @{term "Rt"}, @{typ "t"}, @{thm t_eq}) #> snd
 | 
|  |    303 | *}
 | 
|  |    304 | 
 | 
|  |    305 | section {* lifting of constants *}
 | 
|  |    306 | 
 | 
|  |    307 | text {* information about map-functions for type-constructor *}
 | 
|  |    308 | ML {*
 | 
|  |    309 | type typ_info = {mapfun: string}
 | 
|  |    310 | 
 | 
|  |    311 | local
 | 
|  |    312 |   structure Data = GenericDataFun
 | 
|  |    313 |   (type T = typ_info Symtab.table
 | 
|  |    314 |    val empty = Symtab.empty
 | 
|  |    315 |    val extend = I
 | 
|  |    316 |    fun merge _ = Symtab.merge (K true))
 | 
|  |    317 | in
 | 
|  |    318 |  val lookup = Symtab.lookup o Data.get
 | 
|  |    319 |  fun update k v = Data.map (Symtab.update (k, v)) 
 | 
|  |    320 | end
 | 
|  |    321 | *}
 | 
|  |    322 | 
 | 
|  |    323 | (* mapfuns for some standard types *)
 | 
|  |    324 | setup {*
 | 
|  |    325 |     Context.theory_map (update @{type_name "list"} {mapfun = @{const_name "map"}})
 | 
|  |    326 |  #> Context.theory_map (update @{type_name "*"} {mapfun = @{const_name "prod_fun"}})
 | 
|  |    327 |  #> Context.theory_map (update @{type_name "fun"}  {mapfun = @{const_name "fun_map"}})
 | 
|  |    328 | *}
 | 
|  |    329 | 
 | 
| 2 |    330 | 
 | 
| 0 |    331 | ML {* lookup (Context.Proof @{context}) @{type_name list} *}
 | 
|  |    332 | 
 | 
|  |    333 | ML {*
 | 
|  |    334 | datatype abs_or_rep = abs | rep
 | 
|  |    335 | 
 | 
|  |    336 | fun get_fun abs_or_rep rty qty lthy ty = 
 | 
|  |    337 | let
 | 
|  |    338 |   val qty_name = Long_Name.base_name (fst (dest_Type qty))
 | 
|  |    339 |   
 | 
|  |    340 |   fun get_fun_aux s fs_tys =
 | 
|  |    341 |   let
 | 
|  |    342 |     val (fs, tys) = split_list fs_tys
 | 
|  |    343 |     val (otys, ntys) = split_list tys 
 | 
|  |    344 |     val oty = Type (s, otys)
 | 
|  |    345 |     val nty = Type (s, ntys)
 | 
|  |    346 |     val ftys = map (op -->) tys
 | 
|  |    347 |   in
 | 
|  |    348 |    (case (lookup (Context.Proof lthy) s) of 
 | 
|  |    349 |       SOME info => (list_comb (Const (#mapfun info, ftys ---> oty --> nty), fs), (oty, nty))
 | 
|  |    350 |     | NONE => raise ERROR ("no map association for type " ^ s))
 | 
|  |    351 |   end
 | 
|  |    352 | 
 | 
|  |    353 |   fun get_const abs = (Const ("QuotMain.ABS_" ^ qty_name, rty --> qty), (rty, qty))
 | 
|  |    354 |     | get_const rep = (Const ("QuotMain.REP_" ^ qty_name, qty --> rty), (qty, rty))
 | 
|  |    355 | in
 | 
|  |    356 |   if ty = qty
 | 
|  |    357 |   then (get_const abs_or_rep)
 | 
|  |    358 |   else (case ty of
 | 
|  |    359 |           TFree _ => (Abs ("x", ty, Bound 0), (ty, ty))
 | 
|  |    360 |         | Type (_, []) => (Abs ("x", ty, Bound 0), (ty, ty))
 | 
|  |    361 |         | Type (s, tys) => get_fun_aux s (map (get_fun abs_or_rep rty qty lthy) tys)  
 | 
|  |    362 |         | _ => raise ERROR ("no variables")
 | 
|  |    363 |        )
 | 
|  |    364 | end
 | 
|  |    365 | *}
 | 
|  |    366 | 
 | 
|  |    367 | ML {*
 | 
| 2 |    368 |   get_fun rep @{typ t} @{typ qt} @{context} @{typ "t * nat"}
 | 
|  |    369 |   |> fst
 | 
|  |    370 |   |> Syntax.string_of_term @{context}
 | 
|  |    371 |   |> writeln
 | 
|  |    372 | *}
 | 
|  |    373 | 
 | 
|  |    374 | 
 | 
|  |    375 | ML {*
 | 
| 0 |    376 | fun get_const_def nconst oconst rty qty lthy =
 | 
|  |    377 | let
 | 
|  |    378 |   val ty = fastype_of nconst
 | 
|  |    379 |   val (arg_tys, res_ty) = strip_type ty
 | 
|  |    380 |  
 | 
|  |    381 |   val fresh_args = arg_tys |> map (pair "x")
 | 
|  |    382 |                            |> Variable.variant_frees lthy [nconst, oconst] 
 | 
|  |    383 |                            |> map Free
 | 
|  |    384 | 
 | 
|  |    385 |   val rep_fns = map (fst o get_fun rep rty qty lthy) arg_tys
 | 
|  |    386 |   val abs_fn  = (fst o get_fun abs rty qty lthy) res_ty
 | 
|  |    387 | 
 | 
|  |    388 | in
 | 
|  |    389 |   map (op $) (rep_fns ~~ fresh_args)
 | 
|  |    390 |   |> curry list_comb oconst
 | 
|  |    391 |   |> curry (op $) abs_fn
 | 
|  |    392 |   |> fold_rev lambda fresh_args
 | 
|  |    393 | end
 | 
|  |    394 | *}
 | 
|  |    395 | 
 | 
|  |    396 | ML {*
 | 
|  |    397 | fun exchange_ty rty qty ty = 
 | 
|  |    398 |   if ty = rty then qty
 | 
|  |    399 |   else 
 | 
|  |    400 |     (case ty of
 | 
|  |    401 |        Type (s, tys) => Type (s, map (exchange_ty rty qty) tys)
 | 
|  |    402 |       | _ => ty)
 | 
|  |    403 | *}
 | 
|  |    404 | 
 | 
|  |    405 | ML {*
 | 
|  |    406 | fun make_const_def nconst_name oconst mx rty qty lthy = 
 | 
|  |    407 | let
 | 
|  |    408 |   val oconst_ty = fastype_of oconst
 | 
|  |    409 |   val nconst_ty = exchange_ty rty qty oconst_ty
 | 
|  |    410 |   val nconst = Const (nconst_name, nconst_ty)
 | 
|  |    411 |   val def_trm = get_const_def nconst oconst rty qty lthy
 | 
|  |    412 | in
 | 
|  |    413 |   make_def (Binding.name nconst_name, mx, def_trm) lthy
 | 
|  |    414 | end  
 | 
|  |    415 | *}
 | 
|  |    416 | 
 | 
| 2 |    417 | local_setup {*
 | 
|  |    418 |   make_const_def "VR" @{term "vr"} NoSyn @{typ "t"} @{typ "qt"} #> snd
 | 
|  |    419 | *}
 | 
|  |    420 | 
 | 
|  |    421 | local_setup {*
 | 
|  |    422 |   make_const_def "AP" @{term "ap"} NoSyn @{typ "t"} @{typ "qt"} #> snd
 | 
|  |    423 | *}
 | 
|  |    424 | 
 | 
|  |    425 | local_setup {*
 | 
|  |    426 |   make_const_def "LM" @{term "lm"} NoSyn @{typ "t"} @{typ "qt"} #> snd
 | 
|  |    427 | *}
 | 
|  |    428 | 
 | 
|  |    429 | thm VR_def
 | 
|  |    430 | thm AP_def
 | 
|  |    431 | thm LM_def
 | 
|  |    432 | term LM 
 | 
|  |    433 | term VR
 | 
|  |    434 | term AP
 | 
|  |    435 | 
 | 
|  |    436 | 
 | 
| 0 |    437 | text {* a test with functions *}
 | 
|  |    438 | datatype 'a t' =
 | 
|  |    439 |   vr' "string"
 | 
|  |    440 | | ap' "('a t') * ('a t')"
 | 
|  |    441 | | lm' "'a" "string \<Rightarrow> ('a t')"
 | 
|  |    442 | 
 | 
|  |    443 | consts Rt' :: "('a t') \<Rightarrow> ('a t') \<Rightarrow> bool"
 | 
|  |    444 | axioms t_eq': "EQUIV Rt'"
 | 
|  |    445 | 
 | 
|  |    446 | local_setup {*
 | 
|  |    447 |   typedef_main (@{binding "qt'"}, @{term "Rt'"}, @{typ "'a t'"}, @{thm t_eq'}) #> snd
 | 
|  |    448 | *}
 | 
|  |    449 | 
 | 
| 2 |    450 | 
 | 
| 0 |    451 | local_setup {*
 | 
|  |    452 |   make_const_def "VR'" @{term "vr'"} NoSyn @{typ "'a t'"} @{typ "'a qt'"} #> snd
 | 
|  |    453 | *}
 | 
|  |    454 | 
 | 
|  |    455 | local_setup {*
 | 
|  |    456 |   make_const_def "AP'" @{term "ap'"} NoSyn @{typ "'a t'"} @{typ "'a qt'"} #> snd
 | 
|  |    457 | *}
 | 
|  |    458 | 
 | 
|  |    459 | local_setup {*
 | 
|  |    460 |   make_const_def "LM'" @{term "lm'"} NoSyn @{typ "'a t'"} @{typ "'a qt'"} #> snd
 | 
|  |    461 | *}
 | 
|  |    462 | 
 | 
|  |    463 | thm VR'_def
 | 
|  |    464 | thm AP'_def
 | 
|  |    465 | thm LM'_def
 | 
|  |    466 | term LM' 
 | 
|  |    467 | term VR'
 | 
|  |    468 | term AP'
 | 
|  |    469 | 
 | 
|  |    470 | text {* finite set example *}
 | 
|  |    471 | 
 | 
|  |    472 | inductive 
 | 
|  |    473 |   list_eq ("_ \<approx> _")
 | 
|  |    474 | where
 | 
|  |    475 |   "a#b#xs \<approx> b#a#xs"
 | 
|  |    476 | | "[] \<approx> []"
 | 
|  |    477 | | "xs \<approx> ys \<Longrightarrow> ys \<approx> xs"
 | 
|  |    478 | | "a#a#xs \<approx> a#xs"
 | 
|  |    479 | | "xs \<approx> ys \<Longrightarrow> a#xs \<approx> a#ys"
 | 
|  |    480 | | "\<lbrakk>xs1 \<approx> xs2; xs2 \<approx> xs3\<rbrakk> \<Longrightarrow> xs1 \<approx> xs3"
 | 
|  |    481 | 
 | 
|  |    482 | lemma list_eq_sym:
 | 
|  |    483 |   shows "xs \<approx> xs"
 | 
|  |    484 | apply(induct xs)
 | 
|  |    485 | apply(auto intro: list_eq.intros)
 | 
|  |    486 | done
 | 
|  |    487 | 
 | 
|  |    488 | lemma equiv_list_eq:
 | 
|  |    489 |   shows "EQUIV list_eq"
 | 
|  |    490 | unfolding EQUIV_REFL_SYM_TRANS REFL_def SYM_def TRANS_def
 | 
|  |    491 | apply(auto intro: list_eq.intros list_eq_sym)
 | 
|  |    492 | done
 | 
|  |    493 | 
 | 
|  |    494 | local_setup {*
 | 
|  |    495 |   typedef_main (@{binding "fset"}, @{term "list_eq"}, @{typ "'a list"}, @{thm "equiv_list_eq"}) #> snd
 | 
|  |    496 | *}
 | 
|  |    497 | 
 | 
|  |    498 | typ "'a fset"
 | 
|  |    499 | thm "Rep_fset"
 | 
|  |    500 | 
 | 
|  |    501 | local_setup {*
 | 
|  |    502 |   make_const_def "EMPTY" @{term "[]"} NoSyn @{typ "'a list"} @{typ "'a fset"} #> snd
 | 
|  |    503 | *}
 | 
|  |    504 | 
 | 
|  |    505 | term Nil 
 | 
|  |    506 | term EMPTY
 | 
| 2 |    507 | thm EMPTY_def
 | 
|  |    508 | 
 | 
| 0 |    509 | 
 | 
|  |    510 | local_setup {*
 | 
|  |    511 |   make_const_def "INSERT" @{term "op #"} NoSyn @{typ "'a list"} @{typ "'a fset"} #> snd
 | 
|  |    512 | *}
 | 
|  |    513 | 
 | 
|  |    514 | term Cons 
 | 
|  |    515 | term INSERT
 | 
| 2 |    516 | thm INSERT_def
 | 
| 0 |    517 | 
 | 
|  |    518 | local_setup {*
 | 
|  |    519 |   make_const_def "UNION" @{term "op @"} NoSyn @{typ "'a list"} @{typ "'a fset"} #> snd
 | 
|  |    520 | *}
 | 
|  |    521 | 
 | 
|  |    522 | term append 
 | 
|  |    523 | term UNION
 | 
| 2 |    524 | thm UNION_def
 | 
|  |    525 | 
 | 
| 0 |    526 | thm QUOTIENT_fset
 | 
|  |    527 | 
 | 
|  |    528 | fun 
 | 
|  |    529 |   membship :: "'a \<Rightarrow> 'a list \<Rightarrow> bool" ("_ memb _")
 | 
|  |    530 | where
 | 
|  |    531 |   m1: "(x memb []) = False"
 | 
|  |    532 | | m2: "(x memb (y#xs)) = ((x=y) \<or> (x memb xs))"
 | 
|  |    533 | 
 | 
| 2 |    534 | lemma mem_respects:
 | 
|  |    535 |   fixes z::"nat"
 | 
|  |    536 |   assumes a: "list_eq x y"
 | 
|  |    537 |   shows "z memb x = z memb y"
 | 
|  |    538 | using a
 | 
|  |    539 | apply(induct)
 | 
|  |    540 | apply(auto)
 | 
|  |    541 | done
 | 
|  |    542 | 
 | 
| 0 |    543 | 
 | 
|  |    544 | local_setup {*
 | 
|  |    545 |   make_const_def "IN" @{term "membship"} NoSyn @{typ "'a list"} @{typ "'a fset"} #> snd
 | 
|  |    546 | *}
 | 
|  |    547 | 
 | 
|  |    548 | term membship
 | 
|  |    549 | term IN
 | 
| 2 |    550 | thm IN_def
 | 
| 0 |    551 | 
 | 
| 2 |    552 | lemmas a = QUOT_TYPE.ABS_def[OF QUOT_TYPE_fset]
 | 
|  |    553 | thm QUOT_TYPE.thm11[OF QUOT_TYPE_fset, THEN iffD1, simplified a]
 | 
| 0 |    554 | 
 | 
| 2 |    555 | lemma yy:
 | 
|  |    556 |   shows "(False = x memb []) = (False = IN (x::nat) EMPTY)"
 | 
|  |    557 | unfolding IN_def EMPTY_def
 | 
|  |    558 | apply(rule_tac f="(op =) False" in arg_cong)
 | 
|  |    559 | apply(rule mem_respects)
 | 
|  |    560 | apply(unfold REP_fset_def ABS_fset_def)
 | 
|  |    561 | apply(rule QUOT_TYPE.REP_ABS_rsp[OF QUOT_TYPE_fset])
 | 
|  |    562 | apply(rule list_eq.intros)
 | 
| 0 |    563 | done
 | 
|  |    564 | 
 | 
| 2 |    565 | lemma
 | 
|  |    566 |   shows "IN (x::nat) EMPTY = False"
 | 
|  |    567 | using m1
 | 
|  |    568 | apply -
 | 
|  |    569 | apply(rule yy[THEN iffD1, symmetric])
 | 
|  |    570 | apply(simp)
 | 
| 0 |    571 | done
 | 
|  |    572 | 
 | 
| 2 |    573 | lemma
 | 
|  |    574 |   shows "((x=y) \<or> (IN x xs) = (IN (x::nat) (INSERT y xs))) =
 | 
|  |    575 |          ((x = y) \<or> x memb REP_fset xs = x memb (y # REP_fset xs))"
 | 
|  |    576 | unfolding IN_def INSERT_def
 | 
|  |    577 | apply(rule_tac f="(op \<or>) (x=y)" in arg_cong)
 | 
|  |    578 | apply(rule_tac f="(op =) (x memb REP_fset xs)" in arg_cong)
 | 
|  |    579 | apply(rule mem_respects)
 | 
|  |    580 | apply(rule list_eq.intros(3))
 | 
|  |    581 | apply(unfold REP_fset_def ABS_fset_def)
 | 
|  |    582 | apply(rule QUOT_TYPE.REP_ABS_rsp[OF QUOT_TYPE_fset])
 | 
|  |    583 | apply(rule list_eq_sym)
 | 
|  |    584 | done
 | 
| 0 |    585 | 
 | 
| 3 |    586 | lemma helper:
 | 
|  |    587 |   assumes a : "list_eq l1 r1"
 | 
|  |    588 |   shows "list_eq (l1 @ l) (r1 @ l)"
 | 
|  |    589 |   using a
 | 
|  |    590 |   apply(induct)
 | 
|  |    591 |   apply(simp add:list_eq.intros)
 | 
|  |    592 |   apply(simp add:list_eq_sym)
 | 
|  |    593 |   apply(simp add:list_eq.intros(3))
 | 
|  |    594 |   apply(simp add:list_eq.intros(4))
 | 
|  |    595 |   apply(simp add:list_eq.intros(5))
 | 
|  |    596 |   apply(rule_tac list_eq.intros(6))
 | 
|  |    597 |   apply(assumption)
 | 
|  |    598 |   apply(assumption)
 | 
|  |    599 | done
 | 
|  |    600 | 
 | 
|  |    601 | lemma yyy :
 | 
|  |    602 |   shows "
 | 
|  |    603 |     (
 | 
|  |    604 |      (UNION EMPTY s = s) &
 | 
|  |    605 |      ((UNION (INSERT e s1) s2) = (INSERT e (UNION s1 s2)))
 | 
|  |    606 |     ) = (
 | 
|  |    607 |      ((ABS_fset ([] @ REP_fset s)) = s) &
 | 
|  |    608 |      ((ABS_fset ((e # (REP_fset s1)) @ REP_fset s2)) = ABS_fset (e # (REP_fset s1 @ REP_fset s2)))
 | 
|  |    609 |     )"
 | 
|  |    610 |   unfolding UNION_def EMPTY_def INSERT_def
 | 
|  |    611 |   apply(rule_tac f="(op &)" in arg_cong2)
 | 
|  |    612 |   apply(rule_tac f="(op =)" in arg_cong2)
 | 
|  |    613 |   apply (unfold REP_fset_def ABS_fset_def)
 | 
|  |    614 |   apply(simp only: QUOT_TYPE.thm11[OF QUOT_TYPE_fset,symmetric])
 | 
|  |    615 |   apply(rule helper)
 | 
|  |    616 |   apply(rule list_eq.intros(3))
 | 
|  |    617 |   apply(rule QUOT_TYPE.REP_ABS_rsp[OF QUOT_TYPE_fset])
 | 
|  |    618 |   apply(rule list_eq_sym)
 | 
|  |    619 |   apply(simp)
 | 
|  |    620 |   apply(rule_tac f="(op =)" in arg_cong2)
 | 
|  |    621 |   apply(simp only: QUOT_TYPE.thm11[OF QUOT_TYPE_fset,symmetric])
 | 
|  |    622 |   apply(rule helper)
 | 
|  |    623 |   apply(rule list_eq.intros(3))
 | 
|  |    624 |   apply(rule QUOT_TYPE.REP_ABS_rsp[OF QUOT_TYPE_fset])
 | 
|  |    625 |   apply(rule list_eq_sym)
 | 
|  |    626 |   apply(simp only: QUOT_TYPE.thm11[OF QUOT_TYPE_fset,symmetric])
 | 
|  |    627 |   apply(fold REP_fset_def ABS_fset_def)
 | 
|  |    628 |   apply(rule list_eq.intros(5))
 | 
|  |    629 |   apply(rule list_eq.intros(3))
 | 
|  |    630 |   apply (unfold REP_fset_def ABS_fset_def)
 | 
|  |    631 |   apply(rule QUOT_TYPE.REP_ABS_rsp[OF QUOT_TYPE_fset])
 | 
|  |    632 |   apply(rule list_eq_sym)
 | 
|  |    633 | done
 | 
|  |    634 | 
 | 
|  |    635 | lemma
 | 
|  |    636 |   shows "
 | 
|  |    637 |     (
 | 
|  |    638 |      (UNION EMPTY s = s) &
 | 
|  |    639 |      ((UNION (INSERT e s1) s2) = (INSERT e (UNION s1 s2)))
 | 
|  |    640 |     )"
 | 
|  |    641 |   apply(simp add:yyy)
 | 
|  |    642 |   apply (unfold REP_fset_def ABS_fset_def)
 | 
|  |    643 |   apply (rule QUOT_TYPE.thm10[OF QUOT_TYPE_fset])
 | 
|  |    644 | done
 | 
|  |    645 | 
 |