src/HOL/HOLCF/Tools/Domain/domain_constructors.ML
author wenzelm
Thu, 09 Jun 2011 16:34:49 +0200
changeset 44206 2b47822868e4
parent 43162 b1f544c84040
child 44951 53d95b52954c
permissions -rw-r--r--
discontinued Name.variant to emphasize that this is old-style / indirect;
     1 (*  Title:      HOL/HOLCF/Tools/Domain/domain_constructors.ML
     2     Author:     Brian Huffman
     3 
     4 Defines constructor functions for a given domain isomorphism
     5 and proves related theorems.
     6 *)
     7 
     8 signature DOMAIN_CONSTRUCTORS =
     9 sig
    10   type constr_info =
    11     {
    12       iso_info : Domain_Take_Proofs.iso_info,
    13       con_specs : (term * (bool * typ) list) list,
    14       con_betas : thm list,
    15       nchotomy : thm,
    16       exhaust : thm,
    17       compacts : thm list,
    18       con_rews : thm list,
    19       inverts : thm list,
    20       injects : thm list,
    21       dist_les : thm list,
    22       dist_eqs : thm list,
    23       cases : thm list,
    24       sel_rews : thm list,
    25       dis_rews : thm list,
    26       match_rews : thm list
    27     }
    28   val add_domain_constructors :
    29       binding
    30       -> (binding * (bool * binding option * typ) list * mixfix) list
    31       -> Domain_Take_Proofs.iso_info
    32       -> theory
    33       -> constr_info * theory
    34 end
    35 
    36 
    37 structure Domain_Constructors : DOMAIN_CONSTRUCTORS =
    38 struct
    39 
    40 open HOLCF_Library
    41 
    42 infixr 6 ->>
    43 infix -->>
    44 infix 9 `
    45 
    46 type constr_info =
    47   {
    48     iso_info : Domain_Take_Proofs.iso_info,
    49     con_specs : (term * (bool * typ) list) list,
    50     con_betas : thm list,
    51     nchotomy : thm,
    52     exhaust : thm,
    53     compacts : thm list,
    54     con_rews : thm list,
    55     inverts : thm list,
    56     injects : thm list,
    57     dist_les : thm list,
    58     dist_eqs : thm list,
    59     cases : thm list,
    60     sel_rews : thm list,
    61     dis_rews : thm list,
    62     match_rews : thm list
    63   }
    64 
    65 (************************** miscellaneous functions ***************************)
    66 
    67 val simple_ss = HOL_basic_ss addsimps simp_thms
    68 
    69 val beta_rules =
    70   @{thms beta_cfun cont_id cont_const cont2cont_APP cont2cont_LAM'} @
    71   @{thms cont2cont_fst cont2cont_snd cont2cont_Pair}
    72 
    73 val beta_ss = HOL_basic_ss addsimps (simp_thms @ beta_rules)
    74 
    75 fun define_consts
    76     (specs : (binding * term * mixfix) list)
    77     (thy : theory)
    78     : (term list * thm list) * theory =
    79   let
    80     fun mk_decl (b, t, mx) = (b, fastype_of t, mx)
    81     val decls = map mk_decl specs
    82     val thy = Cont_Consts.add_consts decls thy
    83     fun mk_const (b, T, mx) = Const (Sign.full_name thy b, T)
    84     val consts = map mk_const decls
    85     fun mk_def c (b, t, mx) =
    86       (Binding.suffix_name "_def" b, Logic.mk_equals (c, t))
    87     val defs = map2 mk_def consts specs
    88     val (def_thms, thy) =
    89       Global_Theory.add_defs false (map Thm.no_attributes defs) thy
    90   in
    91     ((consts, def_thms), thy)
    92   end
    93 
    94 fun prove
    95     (thy : theory)
    96     (defs : thm list)
    97     (goal : term)
    98     (tacs : {prems: thm list, context: Proof.context} -> tactic list)
    99     : thm =
   100   let
   101     fun tac {prems, context} =
   102       rewrite_goals_tac defs THEN
   103       EVERY (tacs {prems = map (rewrite_rule defs) prems, context = context})
   104   in
   105     Goal.prove_global thy [] [] goal tac
   106   end
   107 
   108 fun get_vars_avoiding
   109     (taken : string list)
   110     (args : (bool * typ) list)
   111     : (term list * term list) =
   112   let
   113     val Ts = map snd args
   114     val ns = Name.variant_list taken (Datatype_Prop.make_tnames Ts)
   115     val vs = map Free (ns ~~ Ts)
   116     val nonlazy = map snd (filter_out (fst o fst) (args ~~ vs))
   117   in
   118     (vs, nonlazy)
   119   end
   120 
   121 fun get_vars args = get_vars_avoiding [] args
   122 
   123 (************** generating beta reduction rules from definitions **************)
   124 
   125 local
   126   fun arglist (Const _ $ Abs (s, T, t)) =
   127       let
   128         val arg = Free (s, T)
   129         val (args, body) = arglist (subst_bound (arg, t))
   130       in (arg :: args, body) end
   131     | arglist t = ([], t)
   132 in
   133   fun beta_of_def thy def_thm =
   134       let
   135         val (con, lam) = Logic.dest_equals (concl_of def_thm)
   136         val (args, rhs) = arglist lam
   137         val lhs = list_ccomb (con, args)
   138         val goal = mk_equals (lhs, rhs)
   139         val cs = ContProc.cont_thms lam
   140         val betas = map (fn c => mk_meta_eq (c RS @{thm beta_cfun})) cs
   141       in
   142         prove thy (def_thm::betas) goal (K [rtac reflexive_thm 1])
   143       end
   144 end
   145 
   146 (******************************************************************************)
   147 (************* definitions and theorems for constructor functions *************)
   148 (******************************************************************************)
   149 
   150 fun add_constructors
   151     (spec : (binding * (bool * typ) list * mixfix) list)
   152     (abs_const : term)
   153     (iso_locale : thm)
   154     (thy : theory)
   155     =
   156   let
   157 
   158     (* get theorems about rep and abs *)
   159     val abs_strict = iso_locale RS @{thm iso.abs_strict}
   160 
   161     (* get types of type isomorphism *)
   162     val (rhsT, lhsT) = dest_cfunT (fastype_of abs_const)
   163 
   164     fun vars_of args =
   165       let
   166         val Ts = map snd args
   167         val ns = Datatype_Prop.make_tnames Ts
   168       in
   169         map Free (ns ~~ Ts)
   170       end
   171 
   172     (* define constructor functions *)
   173     val ((con_consts, con_defs), thy) =
   174       let
   175         fun one_arg (lazy, T) var = if lazy then mk_up var else var
   176         fun one_con (_,args,_) = mk_stuple (map2 one_arg args (vars_of args))
   177         fun mk_abs t = abs_const ` t
   178         val rhss = map mk_abs (mk_sinjects (map one_con spec))
   179         fun mk_def (bind, args, mx) rhs =
   180           (bind, big_lambdas (vars_of args) rhs, mx)
   181       in
   182         define_consts (map2 mk_def spec rhss) thy
   183       end
   184 
   185     (* prove beta reduction rules for constructors *)
   186     val con_betas = map (beta_of_def thy) con_defs
   187 
   188     (* replace bindings with terms in constructor spec *)
   189     val spec' : (term * (bool * typ) list) list =
   190       let fun one_con con (b, args, mx) = (con, args)
   191       in map2 one_con con_consts spec end
   192 
   193     (* prove exhaustiveness of constructors *)
   194     local
   195       fun arg2typ n (true,  T) = (n+1, mk_upT (TVar (("'a", n), @{sort cpo})))
   196         | arg2typ n (false, T) = (n+1, TVar (("'a", n), @{sort pcpo}))
   197       fun args2typ n [] = (n, oneT)
   198         | args2typ n [arg] = arg2typ n arg
   199         | args2typ n (arg::args) =
   200           let
   201             val (n1, t1) = arg2typ n arg
   202             val (n2, t2) = args2typ n1 args
   203           in (n2, mk_sprodT (t1, t2)) end
   204       fun cons2typ n [] = (n, oneT)
   205         | cons2typ n [con] = args2typ n (snd con)
   206         | cons2typ n (con::cons) =
   207           let
   208             val (n1, t1) = args2typ n (snd con)
   209             val (n2, t2) = cons2typ n1 cons
   210           in (n2, mk_ssumT (t1, t2)) end
   211       val ct = ctyp_of thy (snd (cons2typ 1 spec'))
   212       val thm1 = instantiate' [SOME ct] [] @{thm exh_start}
   213       val thm2 = rewrite_rule (map mk_meta_eq @{thms ex_bottom_iffs}) thm1
   214       val thm3 = rewrite_rule [mk_meta_eq @{thm conj_assoc}] thm2
   215 
   216       val y = Free ("y", lhsT)
   217       fun one_con (con, args) =
   218         let
   219           val (vs, nonlazy) = get_vars_avoiding ["y"] args
   220           val eqn = mk_eq (y, list_ccomb (con, vs))
   221           val conj = foldr1 mk_conj (eqn :: map mk_defined nonlazy)
   222         in Library.foldr mk_ex (vs, conj) end
   223       val goal = mk_trp (foldr1 mk_disj (mk_undef y :: map one_con spec'))
   224       (* first rules replace "y = bottom \/ P" with "rep$y = bottom \/ P" *)
   225       val tacs = [
   226           rtac (iso_locale RS @{thm iso.casedist_rule}) 1,
   227           rewrite_goals_tac [mk_meta_eq (iso_locale RS @{thm iso.iso_swap})],
   228           rtac thm3 1]
   229     in
   230       val nchotomy = prove thy con_betas goal (K tacs)
   231       val exhaust =
   232           (nchotomy RS @{thm exh_casedist0})
   233           |> rewrite_rule @{thms exh_casedists}
   234           |> Drule.zero_var_indexes
   235     end
   236 
   237     (* prove compactness rules for constructors *)
   238     val compacts =
   239       let
   240         val rules = @{thms compact_sinl compact_sinr compact_spair
   241                            compact_up compact_ONE}
   242         val tacs =
   243           [rtac (iso_locale RS @{thm iso.compact_abs}) 1,
   244            REPEAT (resolve_tac rules 1 ORELSE atac 1)]
   245         fun con_compact (con, args) =
   246           let
   247             val vs = vars_of args
   248             val con_app = list_ccomb (con, vs)
   249             val concl = mk_trp (mk_compact con_app)
   250             val assms = map (mk_trp o mk_compact) vs
   251             val goal = Logic.list_implies (assms, concl)
   252           in
   253             prove thy con_betas goal (K tacs)
   254           end
   255       in
   256         map con_compact spec'
   257       end
   258 
   259     (* prove strictness rules for constructors *)
   260     local
   261       fun con_strict (con, args) = 
   262         let
   263           val rules = abs_strict :: @{thms con_strict_rules}
   264           val (vs, nonlazy) = get_vars args
   265           fun one_strict v' =
   266             let
   267               val bottom = mk_bottom (fastype_of v')
   268               val vs' = map (fn v => if v = v' then bottom else v) vs
   269               val goal = mk_trp (mk_undef (list_ccomb (con, vs')))
   270               val tacs = [simp_tac (HOL_basic_ss addsimps rules) 1]
   271             in prove thy con_betas goal (K tacs) end
   272         in map one_strict nonlazy end
   273 
   274       fun con_defin (con, args) =
   275         let
   276           fun iff_disj (t, []) = HOLogic.mk_not t
   277             | iff_disj (t, ts) = mk_eq (t, foldr1 HOLogic.mk_disj ts)
   278           val (vs, nonlazy) = get_vars args
   279           val lhs = mk_undef (list_ccomb (con, vs))
   280           val rhss = map mk_undef nonlazy
   281           val goal = mk_trp (iff_disj (lhs, rhss))
   282           val rule1 = iso_locale RS @{thm iso.abs_bottom_iff}
   283           val rules = rule1 :: @{thms con_bottom_iff_rules}
   284           val tacs = [simp_tac (HOL_ss addsimps rules) 1]
   285         in prove thy con_betas goal (K tacs) end
   286     in
   287       val con_stricts = maps con_strict spec'
   288       val con_defins = map con_defin spec'
   289       val con_rews = con_stricts @ con_defins
   290     end
   291 
   292     (* prove injectiveness of constructors *)
   293     local
   294       fun pgterm rel (con, args) =
   295         let
   296           fun prime (Free (n, T)) = Free (n^"'", T)
   297             | prime t             = t
   298           val (xs, nonlazy) = get_vars args
   299           val ys = map prime xs
   300           val lhs = rel (list_ccomb (con, xs), list_ccomb (con, ys))
   301           val rhs = foldr1 mk_conj (ListPair.map rel (xs, ys))
   302           val concl = mk_trp (mk_eq (lhs, rhs))
   303           val zs = case args of [_] => [] | _ => nonlazy
   304           val assms = map (mk_trp o mk_defined) zs
   305           val goal = Logic.list_implies (assms, concl)
   306         in prove thy con_betas goal end
   307       val cons' = filter (fn (_, args) => not (null args)) spec'
   308     in
   309       val inverts =
   310         let
   311           val abs_below = iso_locale RS @{thm iso.abs_below}
   312           val rules1 = abs_below :: @{thms sinl_below sinr_below spair_below up_below}
   313           val rules2 = @{thms up_defined spair_defined ONE_defined}
   314           val rules = rules1 @ rules2
   315           val tacs = [asm_simp_tac (simple_ss addsimps rules) 1]
   316         in map (fn c => pgterm mk_below c (K tacs)) cons' end
   317       val injects =
   318         let
   319           val abs_eq = iso_locale RS @{thm iso.abs_eq}
   320           val rules1 = abs_eq :: @{thms sinl_eq sinr_eq spair_eq up_eq}
   321           val rules2 = @{thms up_defined spair_defined ONE_defined}
   322           val rules = rules1 @ rules2
   323           val tacs = [asm_simp_tac (simple_ss addsimps rules) 1]
   324         in map (fn c => pgterm mk_eq c (K tacs)) cons' end
   325     end
   326 
   327     (* prove distinctness of constructors *)
   328     local
   329       fun map_dist (f : 'a -> 'a -> 'b) (xs : 'a list) : 'b list =
   330         flat (map_index (fn (i, x) => map (f x) (nth_drop i xs)) xs)
   331       fun prime (Free (n, T)) = Free (n^"'", T)
   332         | prime t             = t
   333       fun iff_disj (t, []) = mk_not t
   334         | iff_disj (t, ts) = mk_eq (t, foldr1 mk_disj ts)
   335       fun iff_disj2 (t, [], us) = mk_not t
   336         | iff_disj2 (t, ts, []) = mk_not t
   337         | iff_disj2 (t, ts, us) =
   338           mk_eq (t, mk_conj (foldr1 mk_disj ts, foldr1 mk_disj us))
   339       fun dist_le (con1, args1) (con2, args2) =
   340         let
   341           val (vs1, zs1) = get_vars args1
   342           val (vs2, zs2) = get_vars args2 |> pairself (map prime)
   343           val lhs = mk_below (list_ccomb (con1, vs1), list_ccomb (con2, vs2))
   344           val rhss = map mk_undef zs1
   345           val goal = mk_trp (iff_disj (lhs, rhss))
   346           val rule1 = iso_locale RS @{thm iso.abs_below}
   347           val rules = rule1 :: @{thms con_below_iff_rules}
   348           val tacs = [simp_tac (HOL_ss addsimps rules) 1]
   349         in prove thy con_betas goal (K tacs) end
   350       fun dist_eq (con1, args1) (con2, args2) =
   351         let
   352           val (vs1, zs1) = get_vars args1
   353           val (vs2, zs2) = get_vars args2 |> pairself (map prime)
   354           val lhs = mk_eq (list_ccomb (con1, vs1), list_ccomb (con2, vs2))
   355           val rhss1 = map mk_undef zs1
   356           val rhss2 = map mk_undef zs2
   357           val goal = mk_trp (iff_disj2 (lhs, rhss1, rhss2))
   358           val rule1 = iso_locale RS @{thm iso.abs_eq}
   359           val rules = rule1 :: @{thms con_eq_iff_rules}
   360           val tacs = [simp_tac (HOL_ss addsimps rules) 1]
   361         in prove thy con_betas goal (K tacs) end
   362     in
   363       val dist_les = map_dist dist_le spec'
   364       val dist_eqs = map_dist dist_eq spec'
   365     end
   366 
   367     val result =
   368       {
   369         con_consts = con_consts,
   370         con_betas = con_betas,
   371         nchotomy = nchotomy,
   372         exhaust = exhaust,
   373         compacts = compacts,
   374         con_rews = con_rews,
   375         inverts = inverts,
   376         injects = injects,
   377         dist_les = dist_les,
   378         dist_eqs = dist_eqs
   379       }
   380   in
   381     (result, thy)
   382   end
   383 
   384 (******************************************************************************)
   385 (**************** definition and theorems for case combinator *****************)
   386 (******************************************************************************)
   387 
   388 fun add_case_combinator
   389     (spec : (term * (bool * typ) list) list)
   390     (lhsT : typ)
   391     (dbind : binding)
   392     (con_betas : thm list)
   393     (exhaust : thm)
   394     (iso_locale : thm)
   395     (rep_const : term)
   396     (thy : theory)
   397     : ((typ -> term) * thm list) * theory =
   398   let
   399 
   400     (* prove rep/abs rules *)
   401     val rep_strict = iso_locale RS @{thm iso.rep_strict}
   402     val abs_inverse = iso_locale RS @{thm iso.abs_iso}
   403 
   404     (* calculate function arguments of case combinator *)
   405     val tns = map fst (Term.add_tfreesT lhsT [])
   406     val resultT = TFree (singleton (Name.variant_list tns) "'t", @{sort pcpo})
   407     fun fTs T = map (fn (_, args) => map snd args -->> T) spec
   408     val fns = Datatype_Prop.indexify_names (map (K "f") spec)
   409     val fs = map Free (fns ~~ fTs resultT)
   410     fun caseT T = fTs T -->> (lhsT ->> T)
   411 
   412     (* definition of case combinator *)
   413     local
   414       val case_bind = Binding.suffix_name "_case" dbind
   415       fun lambda_arg (lazy, v) t =
   416           (if lazy then mk_fup else I) (big_lambda v t)
   417       fun lambda_args []      t = mk_one_case t
   418         | lambda_args (x::[]) t = lambda_arg x t
   419         | lambda_args (x::xs) t = mk_ssplit (lambda_arg x (lambda_args xs t))
   420       fun one_con f (_, args) =
   421         let
   422           val Ts = map snd args
   423           val ns = Name.variant_list fns (Datatype_Prop.make_tnames Ts)
   424           val vs = map Free (ns ~~ Ts)
   425         in
   426           lambda_args (map fst args ~~ vs) (list_ccomb (f, vs))
   427         end
   428       fun mk_sscases [t] = mk_strictify t
   429         | mk_sscases ts = foldr1 mk_sscase ts
   430       val body = mk_sscases (map2 one_con fs spec)
   431       val rhs = big_lambdas fs (mk_cfcomp (body, rep_const))
   432       val ((case_consts, case_defs), thy) =
   433           define_consts [(case_bind, rhs, NoSyn)] thy
   434       val case_name = Sign.full_name thy case_bind
   435     in
   436       val case_def = hd case_defs
   437       fun case_const T = Const (case_name, caseT T)
   438       val case_app = list_ccomb (case_const resultT, fs)
   439       val thy = thy
   440     end
   441 
   442     (* define syntax for case combinator *)
   443     (* TODO: re-implement case syntax using a parse translation *)
   444     local
   445       fun syntax c = Lexicon.mark_const (fst (dest_Const c))
   446       fun xconst c = Long_Name.base_name (fst (dest_Const c))
   447       fun c_ast authentic con = Ast.Constant (if authentic then syntax con else xconst con)
   448       fun showint n = string_of_int (n+1)
   449       fun expvar n = Ast.Variable ("e" ^ showint n)
   450       fun argvar n (m, _) = Ast.Variable ("a" ^ showint n ^ "_" ^ showint m)
   451       fun argvars n args = map_index (argvar n) args
   452       fun app s (l, r) = Ast.mk_appl (Ast.Constant s) [l, r]
   453       val cabs = app "_cabs"
   454       val capp = app @{const_syntax Rep_cfun}
   455       val capps = Library.foldl capp
   456       fun con1 authentic n (con,args) =
   457           Library.foldl capp (c_ast authentic con, argvars n args)
   458       fun case1 authentic (n, c) =
   459           app "_case1" (Ast.strip_positions (con1 authentic n c), expvar n)
   460       fun arg1 (n, (con,args)) = List.foldr cabs (expvar n) (argvars n args)
   461       fun when1 n (m, c) = if n = m then arg1 (n, c) else Ast.Constant @{const_syntax bottom}
   462       val case_constant = Ast.Constant (syntax (case_const dummyT))
   463       fun case_trans authentic =
   464         (if authentic then Syntax.Parse_Print_Rule else Syntax.Parse_Rule)
   465           (app "_case_syntax"
   466             (Ast.Variable "x",
   467              foldr1 (app "_case2") (map_index (case1 authentic) spec)),
   468            capp (capps (case_constant, map_index arg1 spec), Ast.Variable "x"))
   469       fun one_abscon_trans authentic (n, c) =
   470         (if authentic then Syntax.Parse_Print_Rule else Syntax.Parse_Rule)
   471           (cabs (con1 authentic n c, expvar n),
   472            capps (case_constant, map_index (when1 n) spec))
   473       fun abscon_trans authentic =
   474           map_index (one_abscon_trans authentic) spec
   475       val trans_rules : Ast.ast Syntax.trrule list =
   476           case_trans false :: case_trans true ::
   477           abscon_trans false @ abscon_trans true
   478     in
   479       val thy = Sign.add_trrules trans_rules thy
   480     end
   481 
   482     (* prove beta reduction rule for case combinator *)
   483     val case_beta = beta_of_def thy case_def
   484 
   485     (* prove strictness of case combinator *)
   486     val case_strict =
   487       let
   488         val defs = case_beta :: map mk_meta_eq [rep_strict, @{thm cfcomp2}]
   489         val goal = mk_trp (mk_strict case_app)
   490         val rules = @{thms sscase1 ssplit1 strictify1 one_case1}
   491         val tacs = [resolve_tac rules 1]
   492       in prove thy defs goal (K tacs) end
   493         
   494     (* prove rewrites for case combinator *)
   495     local
   496       fun one_case (con, args) f =
   497         let
   498           val (vs, nonlazy) = get_vars args
   499           val assms = map (mk_trp o mk_defined) nonlazy
   500           val lhs = case_app ` list_ccomb (con, vs)
   501           val rhs = list_ccomb (f, vs)
   502           val concl = mk_trp (mk_eq (lhs, rhs))
   503           val goal = Logic.list_implies (assms, concl)
   504           val defs = case_beta :: con_betas
   505           val rules1 = @{thms strictify2 sscase2 sscase3 ssplit2 fup2 ID1}
   506           val rules2 = @{thms con_bottom_iff_rules}
   507           val rules3 = @{thms cfcomp2 one_case2}
   508           val rules = abs_inverse :: rules1 @ rules2 @ rules3
   509           val tacs = [asm_simp_tac (beta_ss addsimps rules) 1]
   510         in prove thy defs goal (K tacs) end
   511     in
   512       val case_apps = map2 one_case spec fs
   513     end
   514 
   515   in
   516     ((case_const, case_strict :: case_apps), thy)
   517   end
   518 
   519 (******************************************************************************)
   520 (************** definitions and theorems for selector functions ***************)
   521 (******************************************************************************)
   522 
   523 fun add_selectors
   524     (spec : (term * (bool * binding option * typ) list) list)
   525     (rep_const : term)
   526     (abs_inv : thm)
   527     (rep_strict : thm)
   528     (rep_bottom_iff : thm)
   529     (con_betas : thm list)
   530     (thy : theory)
   531     : thm list * theory =
   532   let
   533 
   534     (* define selector functions *)
   535     val ((sel_consts, sel_defs), thy) =
   536       let
   537         fun rangeT s = snd (dest_cfunT (fastype_of s))
   538         fun mk_outl s = mk_cfcomp (from_sinl (dest_ssumT (rangeT s)), s)
   539         fun mk_outr s = mk_cfcomp (from_sinr (dest_ssumT (rangeT s)), s)
   540         fun mk_sfst s = mk_cfcomp (sfst_const (dest_sprodT (rangeT s)), s)
   541         fun mk_ssnd s = mk_cfcomp (ssnd_const (dest_sprodT (rangeT s)), s)
   542         fun mk_down s = mk_cfcomp (from_up (dest_upT (rangeT s)), s)
   543 
   544         fun sels_of_arg s (lazy, NONE,   T) = []
   545           | sels_of_arg s (lazy, SOME b, T) =
   546             [(b, if lazy then mk_down s else s, NoSyn)]
   547         fun sels_of_args s [] = []
   548           | sels_of_args s (v :: []) = sels_of_arg s v
   549           | sels_of_args s (v :: vs) =
   550             sels_of_arg (mk_sfst s) v @ sels_of_args (mk_ssnd s) vs
   551         fun sels_of_cons s [] = []
   552           | sels_of_cons s ((con, args) :: []) = sels_of_args s args
   553           | sels_of_cons s ((con, args) :: cs) =
   554             sels_of_args (mk_outl s) args @ sels_of_cons (mk_outr s) cs
   555         val sel_eqns : (binding * term * mixfix) list =
   556             sels_of_cons rep_const spec
   557       in
   558         define_consts sel_eqns thy
   559       end
   560 
   561     (* replace bindings with terms in constructor spec *)
   562     val spec2 : (term * (bool * term option * typ) list) list =
   563       let
   564         fun prep_arg (lazy, NONE, T) sels = ((lazy, NONE, T), sels)
   565           | prep_arg (lazy, SOME _, T) sels =
   566             ((lazy, SOME (hd sels), T), tl sels)
   567         fun prep_con (con, args) sels =
   568             apfst (pair con) (fold_map prep_arg args sels)
   569       in
   570         fst (fold_map prep_con spec sel_consts)
   571       end
   572 
   573     (* prove selector strictness rules *)
   574     val sel_stricts : thm list =
   575       let
   576         val rules = rep_strict :: @{thms sel_strict_rules}
   577         val tacs = [simp_tac (HOL_basic_ss addsimps rules) 1]
   578         fun sel_strict sel =
   579           let
   580             val goal = mk_trp (mk_strict sel)
   581           in
   582             prove thy sel_defs goal (K tacs)
   583           end
   584       in
   585         map sel_strict sel_consts
   586       end
   587 
   588     (* prove selector application rules *)
   589     val sel_apps : thm list =
   590       let
   591         val defs = con_betas @ sel_defs
   592         val rules = abs_inv :: @{thms sel_app_rules}
   593         val tacs = [asm_simp_tac (simple_ss addsimps rules) 1]
   594         fun sel_apps_of (i, (con, args: (bool * term option * typ) list)) =
   595           let
   596             val Ts : typ list = map #3 args
   597             val ns : string list = Datatype_Prop.make_tnames Ts
   598             val vs : term list = map Free (ns ~~ Ts)
   599             val con_app : term = list_ccomb (con, vs)
   600             val vs' : (bool * term) list = map #1 args ~~ vs
   601             fun one_same (n, sel, T) =
   602               let
   603                 val xs = map snd (filter_out fst (nth_drop n vs'))
   604                 val assms = map (mk_trp o mk_defined) xs
   605                 val concl = mk_trp (mk_eq (sel ` con_app, nth vs n))
   606                 val goal = Logic.list_implies (assms, concl)
   607               in
   608                 prove thy defs goal (K tacs)
   609               end
   610             fun one_diff (n, sel, T) =
   611               let
   612                 val goal = mk_trp (mk_eq (sel ` con_app, mk_bottom T))
   613               in
   614                 prove thy defs goal (K tacs)
   615               end
   616             fun one_con (j, (_, args')) : thm list =
   617               let
   618                 fun prep (i, (lazy, NONE, T)) = NONE
   619                   | prep (i, (lazy, SOME sel, T)) = SOME (i, sel, T)
   620                 val sels : (int * term * typ) list =
   621                   map_filter prep (map_index I args')
   622               in
   623                 if i = j
   624                 then map one_same sels
   625                 else map one_diff sels
   626               end
   627           in
   628             flat (map_index one_con spec2)
   629           end
   630       in
   631         flat (map_index sel_apps_of spec2)
   632       end
   633 
   634   (* prove selector definedness rules *)
   635     val sel_defins : thm list =
   636       let
   637         val rules = rep_bottom_iff :: @{thms sel_bottom_iff_rules}
   638         val tacs = [simp_tac (HOL_basic_ss addsimps rules) 1]
   639         fun sel_defin sel =
   640           let
   641             val (T, U) = dest_cfunT (fastype_of sel)
   642             val x = Free ("x", T)
   643             val lhs = mk_eq (sel ` x, mk_bottom U)
   644             val rhs = mk_eq (x, mk_bottom T)
   645             val goal = mk_trp (mk_eq (lhs, rhs))
   646           in
   647             prove thy sel_defs goal (K tacs)
   648           end
   649         fun one_arg (false, SOME sel, T) = SOME (sel_defin sel)
   650           | one_arg _                    = NONE
   651       in
   652         case spec2 of
   653           [(con, args)] => map_filter one_arg args
   654         | _             => []
   655       end
   656 
   657   in
   658     (sel_stricts @ sel_defins @ sel_apps, thy)
   659   end
   660 
   661 (******************************************************************************)
   662 (************ definitions and theorems for discriminator functions ************)
   663 (******************************************************************************)
   664 
   665 fun add_discriminators
   666     (bindings : binding list)
   667     (spec : (term * (bool * typ) list) list)
   668     (lhsT : typ)
   669     (exhaust : thm)
   670     (case_const : typ -> term)
   671     (case_rews : thm list)
   672     (thy : theory) =
   673   let
   674 
   675     fun vars_of args =
   676       let
   677         val Ts = map snd args
   678         val ns = Datatype_Prop.make_tnames Ts
   679       in
   680         map Free (ns ~~ Ts)
   681       end
   682 
   683     (* define discriminator functions *)
   684     local
   685       fun dis_fun i (j, (con, args)) =
   686         let
   687           val (vs, nonlazy) = get_vars args
   688           val tr = if i = j then @{term TT} else @{term FF}
   689         in
   690           big_lambdas vs tr
   691         end
   692       fun dis_eqn (i, bind) : binding * term * mixfix =
   693         let
   694           val dis_bind = Binding.prefix_name "is_" bind
   695           val rhs = list_ccomb (case_const trT, map_index (dis_fun i) spec)
   696         in
   697           (dis_bind, rhs, NoSyn)
   698         end
   699     in
   700       val ((dis_consts, dis_defs), thy) =
   701           define_consts (map_index dis_eqn bindings) thy
   702     end
   703 
   704     (* prove discriminator strictness rules *)
   705     local
   706       fun dis_strict dis =
   707         let val goal = mk_trp (mk_strict dis)
   708         in prove thy dis_defs goal (K [rtac (hd case_rews) 1]) end
   709     in
   710       val dis_stricts = map dis_strict dis_consts
   711     end
   712 
   713     (* prove discriminator/constructor rules *)
   714     local
   715       fun dis_app (i, dis) (j, (con, args)) =
   716         let
   717           val (vs, nonlazy) = get_vars args
   718           val lhs = dis ` list_ccomb (con, vs)
   719           val rhs = if i = j then @{term TT} else @{term FF}
   720           val assms = map (mk_trp o mk_defined) nonlazy
   721           val concl = mk_trp (mk_eq (lhs, rhs))
   722           val goal = Logic.list_implies (assms, concl)
   723           val tacs = [asm_simp_tac (beta_ss addsimps case_rews) 1]
   724         in prove thy dis_defs goal (K tacs) end
   725       fun one_dis (i, dis) =
   726           map_index (dis_app (i, dis)) spec
   727     in
   728       val dis_apps = flat (map_index one_dis dis_consts)
   729     end
   730 
   731     (* prove discriminator definedness rules *)
   732     local
   733       fun dis_defin dis =
   734         let
   735           val x = Free ("x", lhsT)
   736           val simps = dis_apps @ @{thms dist_eq_tr}
   737           val tacs =
   738             [rtac @{thm iffI} 1,
   739              asm_simp_tac (HOL_basic_ss addsimps dis_stricts) 2,
   740              rtac exhaust 1, atac 1,
   741              DETERM_UNTIL_SOLVED (CHANGED
   742                (asm_full_simp_tac (simple_ss addsimps simps) 1))]
   743           val goal = mk_trp (mk_eq (mk_undef (dis ` x), mk_undef x))
   744         in prove thy [] goal (K tacs) end
   745     in
   746       val dis_defins = map dis_defin dis_consts
   747     end
   748 
   749   in
   750     (dis_stricts @ dis_defins @ dis_apps, thy)
   751   end
   752 
   753 (******************************************************************************)
   754 (*************** definitions and theorems for match combinators ***************)
   755 (******************************************************************************)
   756 
   757 fun add_match_combinators
   758     (bindings : binding list)
   759     (spec : (term * (bool * typ) list) list)
   760     (lhsT : typ)
   761     (exhaust : thm)
   762     (case_const : typ -> term)
   763     (case_rews : thm list)
   764     (thy : theory) =
   765   let
   766 
   767     (* get a fresh type variable for the result type *)
   768     val resultT : typ =
   769       let
   770         val ts : string list = map fst (Term.add_tfreesT lhsT [])
   771         val t : string = singleton (Name.variant_list ts) "'t"
   772       in TFree (t, @{sort pcpo}) end
   773 
   774     (* define match combinators *)
   775     local
   776       val x = Free ("x", lhsT)
   777       fun k args = Free ("k", map snd args -->> mk_matchT resultT)
   778       val fail = mk_fail resultT
   779       fun mat_fun i (j, (con, args)) =
   780         let
   781           val (vs, nonlazy) = get_vars_avoiding ["x","k"] args
   782         in
   783           if i = j then k args else big_lambdas vs fail
   784         end
   785       fun mat_eqn (i, (bind, (con, args))) : binding * term * mixfix =
   786         let
   787           val mat_bind = Binding.prefix_name "match_" bind
   788           val funs = map_index (mat_fun i) spec
   789           val body = list_ccomb (case_const (mk_matchT resultT), funs)
   790           val rhs = big_lambda x (big_lambda (k args) (body ` x))
   791         in
   792           (mat_bind, rhs, NoSyn)
   793         end
   794     in
   795       val ((match_consts, match_defs), thy) =
   796           define_consts (map_index mat_eqn (bindings ~~ spec)) thy
   797     end
   798 
   799     (* register match combinators with fixrec package *)
   800     local
   801       val con_names = map (fst o dest_Const o fst) spec
   802       val mat_names = map (fst o dest_Const) match_consts
   803     in
   804       val thy = Fixrec.add_matchers (con_names ~~ mat_names) thy
   805     end
   806 
   807     (* prove strictness of match combinators *)
   808     local
   809       fun match_strict mat =
   810         let
   811           val (T, (U, V)) = apsnd dest_cfunT (dest_cfunT (fastype_of mat))
   812           val k = Free ("k", U)
   813           val goal = mk_trp (mk_eq (mat ` mk_bottom T ` k, mk_bottom V))
   814           val tacs = [asm_simp_tac (beta_ss addsimps case_rews) 1]
   815         in prove thy match_defs goal (K tacs) end
   816     in
   817       val match_stricts = map match_strict match_consts
   818     end
   819 
   820     (* prove match/constructor rules *)
   821     local
   822       val fail = mk_fail resultT
   823       fun match_app (i, mat) (j, (con, args)) =
   824         let
   825           val (vs, nonlazy) = get_vars_avoiding ["k"] args
   826           val (_, (kT, _)) = apsnd dest_cfunT (dest_cfunT (fastype_of mat))
   827           val k = Free ("k", kT)
   828           val lhs = mat ` list_ccomb (con, vs) ` k
   829           val rhs = if i = j then list_ccomb (k, vs) else fail
   830           val assms = map (mk_trp o mk_defined) nonlazy
   831           val concl = mk_trp (mk_eq (lhs, rhs))
   832           val goal = Logic.list_implies (assms, concl)
   833           val tacs = [asm_simp_tac (beta_ss addsimps case_rews) 1]
   834         in prove thy match_defs goal (K tacs) end
   835       fun one_match (i, mat) =
   836           map_index (match_app (i, mat)) spec
   837     in
   838       val match_apps = flat (map_index one_match match_consts)
   839     end
   840 
   841   in
   842     (match_stricts @ match_apps, thy)
   843   end
   844 
   845 (******************************************************************************)
   846 (******************************* main function ********************************)
   847 (******************************************************************************)
   848 
   849 fun add_domain_constructors
   850     (dbind : binding)
   851     (spec : (binding * (bool * binding option * typ) list * mixfix) list)
   852     (iso_info : Domain_Take_Proofs.iso_info)
   853     (thy : theory) =
   854   let
   855     val dname = Binding.name_of dbind
   856     val _ = writeln ("Proving isomorphism properties of domain "^dname^" ...")
   857 
   858     val bindings = map #1 spec
   859 
   860     (* retrieve facts about rep/abs *)
   861     val lhsT = #absT iso_info
   862     val {rep_const, abs_const, ...} = iso_info
   863     val abs_iso_thm = #abs_inverse iso_info
   864     val rep_iso_thm = #rep_inverse iso_info
   865     val iso_locale = @{thm iso.intro} OF [abs_iso_thm, rep_iso_thm]
   866     val rep_strict = iso_locale RS @{thm iso.rep_strict}
   867     val abs_strict = iso_locale RS @{thm iso.abs_strict}
   868     val rep_bottom_iff = iso_locale RS @{thm iso.rep_bottom_iff}
   869     val abs_bottom_iff = iso_locale RS @{thm iso.abs_bottom_iff}
   870     val iso_rews = [abs_iso_thm, rep_iso_thm, abs_strict, rep_strict]
   871 
   872     (* qualify constants and theorems with domain name *)
   873     val thy = Sign.add_path dname thy
   874 
   875     (* define constructor functions *)
   876     val (con_result, thy) =
   877       let
   878         fun prep_arg (lazy, sel, T) = (lazy, T)
   879         fun prep_con (b, args, mx) = (b, map prep_arg args, mx)
   880         val con_spec = map prep_con spec
   881       in
   882         add_constructors con_spec abs_const iso_locale thy
   883       end
   884     val {con_consts, con_betas, nchotomy, exhaust, compacts, con_rews,
   885           inverts, injects, dist_les, dist_eqs} = con_result
   886 
   887     (* prepare constructor spec *)
   888     val con_specs : (term * (bool * typ) list) list =
   889       let
   890         fun prep_arg (lazy, sel, T) = (lazy, T)
   891         fun prep_con c (b, args, mx) = (c, map prep_arg args)
   892       in
   893         map2 prep_con con_consts spec
   894       end
   895 
   896     (* define case combinator *)
   897     val ((case_const : typ -> term, cases : thm list), thy) =
   898         add_case_combinator con_specs lhsT dbind
   899           con_betas exhaust iso_locale rep_const thy
   900 
   901     (* define and prove theorems for selector functions *)
   902     val (sel_thms : thm list, thy : theory) =
   903       let
   904         val sel_spec : (term * (bool * binding option * typ) list) list =
   905           map2 (fn con => fn (b, args, mx) => (con, args)) con_consts spec
   906       in
   907         add_selectors sel_spec rep_const
   908           abs_iso_thm rep_strict rep_bottom_iff con_betas thy
   909       end
   910 
   911     (* define and prove theorems for discriminator functions *)
   912     val (dis_thms : thm list, thy : theory) =
   913         add_discriminators bindings con_specs lhsT
   914           exhaust case_const cases thy
   915 
   916     (* define and prove theorems for match combinators *)
   917     val (match_thms : thm list, thy : theory) =
   918         add_match_combinators bindings con_specs lhsT
   919           exhaust case_const cases thy
   920 
   921     (* restore original signature path *)
   922     val thy = Sign.parent_path thy
   923 
   924     (* bind theorem names in global theory *)
   925     val (_, thy) =
   926       let
   927         fun qualified name = Binding.qualified true name dbind
   928         val names = "bottom" :: map (fn (b,_,_) => Binding.name_of b) spec
   929         val dname = fst (dest_Type lhsT)
   930         val simp = Simplifier.simp_add
   931         val case_names = Rule_Cases.case_names names
   932         val cases_type = Induct.cases_type dname
   933       in
   934         Global_Theory.add_thmss [
   935           ((qualified "iso_rews"  , iso_rews    ), [simp]),
   936           ((qualified "nchotomy"  , [nchotomy]  ), []),
   937           ((qualified "exhaust"   , [exhaust]   ), [case_names, cases_type]),
   938           ((qualified "case_rews" , cases       ), [simp]),
   939           ((qualified "compacts"  , compacts    ), [simp]),
   940           ((qualified "con_rews"  , con_rews    ), [simp]),
   941           ((qualified "sel_rews"  , sel_thms    ), [simp]),
   942           ((qualified "dis_rews"  , dis_thms    ), [simp]),
   943           ((qualified "dist_les"  , dist_les    ), [simp]),
   944           ((qualified "dist_eqs"  , dist_eqs    ), [simp]),
   945           ((qualified "inverts"   , inverts     ), [simp]),
   946           ((qualified "injects"   , injects     ), [simp]),
   947           ((qualified "match_rews", match_thms  ), [simp])] thy
   948       end
   949 
   950     val result =
   951       {
   952         iso_info = iso_info,
   953         con_specs = con_specs,
   954         con_betas = con_betas,
   955         nchotomy = nchotomy,
   956         exhaust = exhaust,
   957         compacts = compacts,
   958         con_rews = con_rews,
   959         inverts = inverts,
   960         injects = injects,
   961         dist_les = dist_les,
   962         dist_eqs = dist_eqs,
   963         cases = cases,
   964         sel_rews = sel_thms,
   965         dis_rews = dis_thms,
   966         match_rews = match_thms
   967       }
   968   in
   969     (result, thy)
   970   end
   971 
   972 end