QuotMain.thy
changeset 41 72d63aa8af68
parent 40 c8187c293587
child 44 f078dbf557b7
equal deleted inserted replaced
40:c8187c293587 41:72d63aa8af68
   152 in
   152 in
   153   lambda c
   153   lambda c
   154     (HOLogic.exists_const ty $ 
   154     (HOLogic.exists_const ty $ 
   155        lambda x (HOLogic.mk_eq (c, (rel $ x))))
   155        lambda x (HOLogic.mk_eq (c, (rel $ x))))
   156 end
   156 end
   157 *}
   157 
   158 
   158 
   159 ML {*
       
   160 (* makes the new type definitions and proves non-emptyness*)
   159 (* makes the new type definitions and proves non-emptyness*)
   161 fun typedef_make (qty_name, rel, ty) lthy =
   160 fun typedef_make (qty_name, rel, ty) lthy =
   162 let  
   161 let  
   163   val typedef_tac =
   162   val typedef_tac =
   164      EVERY1 [rewrite_goal_tac @{thms mem_def},
   163      EVERY1 [rewrite_goal_tac @{thms mem_def},
   170     (Typedef.add_typedef false NONE
   169     (Typedef.add_typedef false NONE
   171        (qty_name, map fst (Term.add_tfreesT ty []), NoSyn)
   170        (qty_name, map fst (Term.add_tfreesT ty []), NoSyn)
   172          (typedef_term rel ty lthy)
   171          (typedef_term rel ty lthy)
   173            NONE typedef_tac) lthy
   172            NONE typedef_tac) lthy
   174 end
   173 end
   175 *}
   174 
   176 
   175 
   177 ML {*
       
   178 (* tactic to prove the QUOT_TYPE theorem for the new type *)
   176 (* tactic to prove the QUOT_TYPE theorem for the new type *)
   179 fun typedef_quot_type_tac equiv_thm (typedef_info: Typedef.info) =
   177 fun typedef_quot_type_tac equiv_thm (typedef_info: Typedef.info) =
   180 let
   178 let
   181   val rep_thm = #Rep typedef_info
   179   val rep_thm = #Rep typedef_info
   182   val rep_inv = #Rep_inverse typedef_info
   180   val rep_inv = #Rep_inverse typedef_info
   193           rtac rep_inv,
   191           rtac rep_inv,
   194           rtac abs_inv_simpd, rtac @{thm exI}, rtac @{thm refl},
   192           rtac abs_inv_simpd, rtac @{thm exI}, rtac @{thm refl},
   195           rtac rep_inj]
   193           rtac rep_inj]
   196 end
   194 end
   197 
   195 
       
   196 
   198 fun typedef_quot_type_thm (rel, abs, rep, equiv_thm, typedef_info) lthy =
   197 fun typedef_quot_type_thm (rel, abs, rep, equiv_thm, typedef_info) lthy =
   199 let
   198 let
   200   val quot_type_const = Const (@{const_name "QUOT_TYPE"}, dummyT)
   199   val quot_type_const = Const (@{const_name "QUOT_TYPE"}, dummyT)
   201   val goal = HOLogic.mk_Trueprop (quot_type_const $ rel $ abs $ rep)
   200   val goal = HOLogic.mk_Trueprop (quot_type_const $ rel $ abs $ rep)
   202              |> Syntax.check_term lthy
   201              |> Syntax.check_term lthy
   203 in
   202 in
   204   Goal.prove lthy [] [] goal
   203   Goal.prove lthy [] [] goal
   205     (K (typedef_quot_type_tac equiv_thm typedef_info))
   204     (K (typedef_quot_type_tac equiv_thm typedef_info))
   206 end
   205 end
   207 *}
   206 
   208 
   207 
   209 ML {*
       
   210 (* proves the quotient theorem *)
   208 (* proves the quotient theorem *)
   211 fun typedef_quotient_thm (rel, abs, rep, abs_def, rep_def, quot_type_thm) lthy =
   209 fun typedef_quotient_thm (rel, abs, rep, abs_def, rep_def, quot_type_thm) lthy =
   212 let
   210 let
   213   val quotient_const = Const (@{const_name "QUOTIENT"}, dummyT)
   211   val quotient_const = Const (@{const_name "QUOTIENT"}, dummyT)
   214   val goal = HOLogic.mk_Trueprop (quotient_const $ rel $ abs $ rep)
   212   val goal = HOLogic.mk_Trueprop (quotient_const $ rel $ abs $ rep)
   428     val nty = Type (s, ntys)
   426     val nty = Type (s, ntys)
   429     val ftys = map (op -->) tys
   427     val ftys = map (op -->) tys
   430   in
   428   in
   431    (case (lookup (Context.Proof lthy) s) of
   429    (case (lookup (Context.Proof lthy) s) of
   432       SOME info => (list_comb (Const (#mapfun info, ftys ---> oty --> nty), fs), (oty, nty))
   430       SOME info => (list_comb (Const (#mapfun info, ftys ---> oty --> nty), fs), (oty, nty))
   433     | NONE => raise ERROR ("no map association for type " ^ s))
   431     | NONE      => raise ERROR ("no map association for type " ^ s))
   434   end
   432   end
   435 
   433 
   436   fun get_const abs = (Const ("QuotMain.ABS_" ^ qty_name, rty --> qty), (rty, qty))
   434   fun get_const abs = (Const ("QuotMain.ABS_" ^ qty_name, rty --> qty), (rty, qty))
   437     | get_const rep = (Const ("QuotMain.REP_" ^ qty_name, qty --> rty), (qty, rty))
   435     | get_const rep = (Const ("QuotMain.REP_" ^ qty_name, qty --> rty), (qty, rty))
       
   436 
       
   437   fun mk_identity ty = Abs ("x", ty, Bound 0)
       
   438 
   438 in
   439 in
   439   if ty = qty
   440   if ty = qty
   440   then (get_const abs_or_rep)
   441   then (get_const abs_or_rep)
   441   else (case ty of
   442   else (case ty of
   442           TFree _ => (Abs ("x", ty, Bound 0), (ty, ty))
   443           TFree _ => (mk_identity ty, (ty, ty))
   443         | Type (_, []) => (Abs ("x", ty, Bound 0), (ty, ty))
   444         | Type (_, []) => (mk_identity ty, (ty, ty))
   444         | Type (s, tys) => get_fun_aux s (map (get_fun abs_or_rep rty qty lthy) tys)
   445         | Type (s, tys) => get_fun_aux s (map (get_fun abs_or_rep rty qty lthy) tys)
   445         | _ => raise ERROR ("no variables")
   446         | _ => raise ERROR ("no type variables")
   446        )
   447        )
   447 end
   448 end
   448 *}
   449 *}
   449 
   450 
   450 ML {*
   451 ML {*
   453   |> Syntax.string_of_term @{context}
   454   |> Syntax.string_of_term @{context}
   454   |> writeln
   455   |> writeln
   455 *}
   456 *}
   456 
   457 
   457 
   458 
       
   459 text {* produces the definition for a lifted constant *}
   458 ML {*
   460 ML {*
   459 fun get_const_def nconst oconst rty qty lthy =
   461 fun get_const_def nconst oconst rty qty lthy =
   460 let
   462 let
   461   val ty = fastype_of nconst
   463   val ty = fastype_of nconst
   462   val (arg_tys, res_ty) = strip_type ty
   464   val (arg_tys, res_ty) = strip_type ty
   476 end
   478 end
   477 *}
   479 *}
   478 
   480 
   479 ML {*
   481 ML {*
   480 fun exchange_ty rty qty ty =
   482 fun exchange_ty rty qty ty =
   481   if ty = rty then qty
   483   if ty = rty 
       
   484   then qty
   482   else
   485   else
   483     (case ty of
   486     (case ty of
   484        Type (s, tys) => Type (s, map (exchange_ty rty qty) tys)
   487        Type (s, tys) => Type (s, map (exchange_ty rty qty) tys)
   485       | _ => ty)
   488       | _ => ty
       
   489     )
   486 *}
   490 *}
   487 
   491 
   488 ML {*
   492 ML {*
   489 fun make_const_def nconst_bname oconst mx rty qty lthy =
   493 fun make_const_def nconst_bname oconst mx rty qty lthy =
   490 let
   494 let
   496   make_def (nconst_bname, mx, def_trm) lthy
   500   make_def (nconst_bname, mx, def_trm) lthy
   497 end
   501 end
   498 *}
   502 *}
   499 
   503 
   500 local_setup {*
   504 local_setup {*
   501   make_const_def @{binding VR} @{term "vr"} NoSyn @{typ "t"} @{typ "qt"} #> snd
   505   make_const_def @{binding VR} @{term "vr"} NoSyn @{typ "t"} @{typ "qt"} #> snd #>
   502 *}
   506   make_const_def @{binding AP} @{term "ap"} NoSyn @{typ "t"} @{typ "qt"} #> snd #>
   503 
       
   504 local_setup {*
       
   505   make_const_def @{binding AP} @{term "ap"} NoSyn @{typ "t"} @{typ "qt"} #> snd
       
   506 *}
       
   507 
       
   508 local_setup {*
       
   509   make_const_def @{binding LM} @{term "lm"} NoSyn @{typ "t"} @{typ "qt"} #> snd
   507   make_const_def @{binding LM} @{term "lm"} NoSyn @{typ "t"} @{typ "qt"} #> snd
   510 *}
   508 *}
   511 
   509 
   512 thm VR_def
   510 term vr
   513 thm AP_def
   511 term ap
   514 thm LM_def
   512 term lm
       
   513 thm VR_def AP_def LM_def
   515 term LM
   514 term LM
   516 term VR
   515 term VR
   517 term AP
   516 term AP
   518 
   517 
   519 
   518 
   532 *}
   531 *}
   533 
   532 
   534 print_theorems
   533 print_theorems
   535 
   534 
   536 local_setup {*
   535 local_setup {*
   537   make_const_def @{binding VR'} @{term "vr'"} NoSyn @{typ "'a t'"} @{typ "'a qt'"} #> snd
   536   make_const_def @{binding VR'} @{term "vr'"} NoSyn @{typ "'a t'"} @{typ "'a qt'"} #> snd #>
   538 *}
   537   make_const_def @{binding AP'} @{term "ap'"} NoSyn @{typ "'a t'"} @{typ "'a qt'"} #> snd #>
   539 
       
   540 local_setup {*
       
   541   make_const_def @{binding AP'} @{term "ap'"} NoSyn @{typ "'a t'"} @{typ "'a qt'"} #> snd
       
   542 *}
       
   543 
       
   544 local_setup {*
       
   545   make_const_def @{binding LM'} @{term "lm'"} NoSyn @{typ "'a t'"} @{typ "'a qt'"} #> snd
   538   make_const_def @{binding LM'} @{term "lm'"} NoSyn @{typ "'a t'"} @{typ "'a qt'"} #> snd
   546 *}
   539 *}
   547 
   540 
   548 thm VR'_def
   541 term vr'
   549 thm AP'_def
   542 term ap'
   550 thm LM'_def
   543 term ap'
       
   544 thm VR'_def AP'_def LM'_def
   551 term LM'
   545 term LM'
   552 term VR'
   546 term VR'
   553 term AP'
   547 term AP'
   554 
   548 
   555 
   549 
   646 
   640 
   647 (* text {*
   641 (* text {*
   648  Maybe make_const_def should require a theorem that says that the particular lifted function
   642  Maybe make_const_def should require a theorem that says that the particular lifted function
   649  respects the relation. With it such a definition would be impossible:
   643  respects the relation. With it such a definition would be impossible:
   650  make_const_def @{binding CARD} @{term "length"} NoSyn @{typ "'a list"} @{typ "'a fset"} #> snd
   644  make_const_def @{binding CARD} @{term "length"} NoSyn @{typ "'a list"} @{typ "'a fset"} #> snd
   651 *} *)
   645 *}*)
   652 
   646 
   653 lemma card1_rsp:
   647 lemma card1_rsp:
   654   fixes a b :: "'a list"
   648   fixes a b :: "'a list"
   655   assumes e: "a \<approx> b"
   649   assumes e: "a \<approx> b"
   656   shows "card1 a = card1 b"
   650   shows "card1 a = card1 b"
   821   apply (simp add: QUOT_TYPE_I_fset.thm10)
   815   apply (simp add: QUOT_TYPE_I_fset.thm10)
   822   done
   816   done
   823 
   817 
   824 ML {*
   818 ML {*
   825   fun mk_rep_abs x = @{term REP_fset} $ (@{term ABS_fset} $ x)
   819   fun mk_rep_abs x = @{term REP_fset} $ (@{term ABS_fset} $ x)
   826   val consts = [@{const_name "Nil"}, @{const_name "append"}, @{const_name "Cons"}, @{const_name "membship"}, @{const_name "card1"}]
   820   val consts = [@{const_name "Nil"}, @{const_name "append"}, 
   827 *}
   821                 @{const_name "Cons"}, @{const_name "membship"}, 
   828 
   822                 @{const_name "card1"}]
   829 ML {*
   823 *}
   830 fun build_goal thm constructors lifted_type mk_rep_abs =
   824 
   831   let
   825 ML {*
       
   826 fun build_goal ctxt thm constructors qty mk_rep_abs =
       
   827 let
   832     fun is_const (Const (x, t)) = x mem constructors
   828     fun is_const (Const (x, t)) = x mem constructors
   833       | is_const _ = false
   829       | is_const _ = false
       
   830   
   834     fun maybe_mk_rep_abs t =
   831     fun maybe_mk_rep_abs t =
   835       let
   832     let
   836         val _ = writeln ("Maybe: " ^ Syntax.string_of_term @{context} t)
   833       val _ = writeln ("Maybe: " ^ Syntax.string_of_term ctxt t)
   837       in
   834     in
   838         if type_of t = lifted_type then mk_rep_abs t else t
   835       if type_of t = qty then mk_rep_abs t else t
   839       end
   836     end
   840       handle TYPE _ => t
   837     handle TYPE _ => t
       
   838   
   841     fun build_aux (Abs (s, t, tr)) = (Abs (s, t, build_aux tr))
   839     fun build_aux (Abs (s, t, tr)) = (Abs (s, t, build_aux tr))
   842       | build_aux (f $ a) =
   840       | build_aux (f $ a) =
   843           let
   841           let
   844             val (f, args) = strip_comb (f $ a)
   842             val (f, args) = strip_comb (f $ a)
   845             val _ = writeln (Syntax.string_of_term @{context} f)
   843             val _ = writeln (Syntax.string_of_term ctxt f)
   846            in
   844            in
   847             (if is_const f then maybe_mk_rep_abs (list_comb (f, (map maybe_mk_rep_abs (map build_aux args))))
   845             (if is_const f then maybe_mk_rep_abs (list_comb (f, (map maybe_mk_rep_abs (map build_aux args))))
   848             else list_comb ((build_aux f), (map build_aux args)))
   846             else list_comb ((build_aux f), (map build_aux args)))
   849           end
   847           end
   850       | build_aux x =
   848       | build_aux x =
   851           if is_const x then maybe_mk_rep_abs x else x
   849           if is_const x then maybe_mk_rep_abs x else x
       
   850     
   852     val concl2 = term_of (#prop (crep_thm thm))
   851     val concl2 = term_of (#prop (crep_thm thm))
   853   in
   852 in
   854   Logic.mk_equals ((build_aux concl2), concl2)
   853   Logic.mk_equals ((build_aux concl2), concl2)
   855 end *}
   854 end *}
   856 
   855 
   857 thm EMPTY_def IN_def UNION_def
   856 thm EMPTY_def IN_def UNION_def
   858 
   857 
   940     ])
   939     ])
   941 *}
   940 *}
   942 
   941 
   943 ML {*
   942 ML {*
   944   val m1_novars = snd(no_vars ((Context.Theory @{theory}),@{thm m1}))
   943   val m1_novars = snd(no_vars ((Context.Theory @{theory}),@{thm m1}))
   945   val goal = build_goal m1_novars consts @{typ "'a list"} mk_rep_abs
   944   val goal = build_goal @{context} m1_novars consts @{typ "'a list"} mk_rep_abs
   946   val cgoal = cterm_of @{theory} goal
   945   val cgoal = cterm_of @{theory} goal
   947   val cgoal2 = Thm.rhs_of (MetaSimplifier.rewrite false fset_defs_sym cgoal)
   946   val cgoal2 = Thm.rhs_of (MetaSimplifier.rewrite false fset_defs_sym cgoal)
   948 *}
   947 *}
   949 
   948 
   950 (*notation ( output) "prop" ("#_" [1000] 1000) *)
   949 (*notation ( output) "prop" ("#_" [1000] 1000) *)
   956   done
   955   done
   957 
   956 
   958 thm length_append (* Not true but worth checking that the goal is correct *)
   957 thm length_append (* Not true but worth checking that the goal is correct *)
   959 ML {*
   958 ML {*
   960   val m1_novars = snd(no_vars ((Context.Theory @{theory}),@{thm length_append}))
   959   val m1_novars = snd(no_vars ((Context.Theory @{theory}),@{thm length_append}))
   961   val goal = build_goal m1_novars consts @{typ "'a list"} mk_rep_abs
   960   val goal = build_goal @{context} m1_novars consts @{typ "'a list"} mk_rep_abs
   962   val cgoal = cterm_of @{theory} goal
   961   val cgoal = cterm_of @{theory} goal
   963   val cgoal2 = Thm.rhs_of (MetaSimplifier.rewrite false fset_defs_sym cgoal)
   962   val cgoal2 = Thm.rhs_of (MetaSimplifier.rewrite false fset_defs_sym cgoal)
   964 *}
   963 *}
   965 prove {* Thm.term_of cgoal2 *}
   964 prove {* Thm.term_of cgoal2 *}
   966   apply (tactic {* LocalDefs.unfold_tac @{context} fset_defs *} )
   965   apply (tactic {* LocalDefs.unfold_tac @{context} fset_defs *} )
   968   sorry
   967   sorry
   969 
   968 
   970 thm m2
   969 thm m2
   971 ML {*
   970 ML {*
   972   val m1_novars = snd(no_vars ((Context.Theory @{theory}),@{thm m2}))
   971   val m1_novars = snd(no_vars ((Context.Theory @{theory}),@{thm m2}))
   973   val goal = build_goal m1_novars consts @{typ "'a list"} mk_rep_abs
   972   val goal = build_goal @{context} m1_novars consts @{typ "'a list"} mk_rep_abs
   974   val cgoal = cterm_of @{theory} goal
   973   val cgoal = cterm_of @{theory} goal
   975   val cgoal2 = Thm.rhs_of (MetaSimplifier.rewrite false fset_defs_sym cgoal)
   974   val cgoal2 = Thm.rhs_of (MetaSimplifier.rewrite false fset_defs_sym cgoal)
   976 *}
   975 *}
   977 prove {* Thm.term_of cgoal2 *}
   976 prove {* Thm.term_of cgoal2 *}
   978   apply (tactic {* LocalDefs.unfold_tac @{context} fset_defs *} )
   977   apply (tactic {* LocalDefs.unfold_tac @{context} fset_defs *} )
   980   done
   979   done
   981 
   980 
   982 thm list_eq.intros(4)
   981 thm list_eq.intros(4)
   983 ML {*
   982 ML {*
   984   val m1_novars = snd(no_vars ((Context.Theory @{theory}),@{thm list_eq.intros(4)}))
   983   val m1_novars = snd(no_vars ((Context.Theory @{theory}),@{thm list_eq.intros(4)}))
   985   val goal = build_goal m1_novars consts @{typ "'a list"} mk_rep_abs
   984   val goal = build_goal @{context} m1_novars consts @{typ "'a list"} mk_rep_abs
   986   val cgoal = cterm_of @{theory} goal
   985   val cgoal = cterm_of @{theory} goal
   987   val cgoal2 = Thm.rhs_of (MetaSimplifier.rewrite false fset_defs_sym cgoal)
   986   val cgoal2 = Thm.rhs_of (MetaSimplifier.rewrite false fset_defs_sym cgoal)
   988   val cgoal3 = Thm.rhs_of (MetaSimplifier.rewrite true @{thms QUOT_TYPE_I_fset.thm10} cgoal2)
   987   val cgoal3 = Thm.rhs_of (MetaSimplifier.rewrite true @{thms QUOT_TYPE_I_fset.thm10} cgoal2)
   989 *}
   988 *}
   990 
   989 
  1022 
  1021 
  1023 
  1022 
  1024 thm list_eq.intros(5)
  1023 thm list_eq.intros(5)
  1025 ML {*
  1024 ML {*
  1026   val m1_novars = snd(no_vars ((Context.Theory @{theory}),@{thm list_eq.intros(5)}))
  1025   val m1_novars = snd(no_vars ((Context.Theory @{theory}),@{thm list_eq.intros(5)}))
  1027   val goal = build_goal m1_novars consts @{typ "'a list"} mk_rep_abs
  1026   val goal = build_goal @{context} m1_novars consts @{typ "'a list"} mk_rep_abs
  1028 *}
  1027 *}
  1029 ML {*
  1028 ML {*
  1030   val cgoal = 
  1029   val cgoal = 
  1031     Toplevel.program (fn () =>
  1030     Toplevel.program (fn () =>
  1032       cterm_of @{theory} goal
  1031       cterm_of @{theory} goal
  1047   val m1_novars = snd(no_vars ((Context.Theory @{theory}),@{thm list.induct}))
  1046   val m1_novars = snd(no_vars ((Context.Theory @{theory}),@{thm list.induct}))
  1048 *}
  1047 *}
  1049 ML {*
  1048 ML {*
  1050   val goal =
  1049   val goal =
  1051     Toplevel.program (fn () =>
  1050     Toplevel.program (fn () =>
  1052       build_goal m1_novars consts @{typ "'a list"} mk_rep_abs
  1051       build_goal @{context} m1_novars consts @{typ "'a list"} mk_rep_abs
  1053     )
  1052     )
  1054 *}
  1053 *}
  1055 ML {*
  1054 ML {*
  1056   val cgoal = cterm_of @{theory} goal
  1055   val cgoal = cterm_of @{theory} goal
  1057   val cgoal2 = Thm.rhs_of (MetaSimplifier.rewrite false fset_defs_sym cgoal)
  1056   val cgoal2 = Thm.rhs_of (MetaSimplifier.rewrite false fset_defs_sym cgoal)
  1070 
  1069 
  1071 ML {*
  1070 ML {*
  1072   val m1_novars = snd(no_vars ((Context.Theory @{theory}),@{thm card1_suc}))
  1071   val m1_novars = snd(no_vars ((Context.Theory @{theory}),@{thm card1_suc}))
  1073 *}
  1072 *}
  1074 ML {*
  1073 ML {*
  1075   val goal = build_goal m1_novars consts @{typ "'a list"} mk_rep_abs
  1074   val goal = build_goal @{context} m1_novars consts @{typ "'a list"} mk_rep_abs
  1076 *}
  1075 *}
  1077 ML {*
  1076 ML {*
  1078   val cgoal = cterm_of @{theory} goal
  1077   val cgoal = cterm_of @{theory} goal
  1079   val cgoal2 = Thm.rhs_of (MetaSimplifier.rewrite false fset_defs_sym cgoal)
  1078   val cgoal2 = Thm.rhs_of (MetaSimplifier.rewrite false fset_defs_sym cgoal)
  1080 *}
  1079 *}