rewrite domain package code for selector functions
authorhuffman
Thu, 25 Feb 2010 13:16:28 -0800
changeset 35446 b719dad322fa
parent 35445 593c184977a4
child 35447 82af95d998e0
rewrite domain package code for selector functions
src/HOLCF/Domain.thy
src/HOLCF/Tools/Domain/domain_axioms.ML
src/HOLCF/Tools/Domain/domain_constructors.ML
src/HOLCF/Tools/Domain/domain_syntax.ML
src/HOLCF/Tools/Domain/domain_theorems.ML
--- a/src/HOLCF/Domain.thy	Wed Feb 24 17:37:59 2010 -0800
+++ b/src/HOLCF/Domain.thy	Thu Feb 25 13:16:28 2010 -0800
@@ -229,6 +229,24 @@
 lemmas con_eq_iff_rules =
   sinl_eq sinr_eq sinl_eq_sinr sinr_eq_sinl con_defined_iff_rules
 
+lemmas sel_strict_rules =
+  cfcomp2 sscase1 sfst_strict ssnd_strict fup1
+
+lemma sel_app_extra_rules:
+  "sscase\<cdot>ID\<cdot>\<bottom>\<cdot>(sinr\<cdot>x) = \<bottom>"
+  "sscase\<cdot>ID\<cdot>\<bottom>\<cdot>(sinl\<cdot>x) = x"
+  "sscase\<cdot>\<bottom>\<cdot>ID\<cdot>(sinl\<cdot>x) = \<bottom>"
+  "sscase\<cdot>\<bottom>\<cdot>ID\<cdot>(sinr\<cdot>x) = x"
+  "fup\<cdot>ID\<cdot>(up\<cdot>x) = x"
+by (cases "x = \<bottom>", simp, simp)+
+
+lemmas sel_app_rules =
+  sel_strict_rules sel_app_extra_rules
+  ssnd_spair sfst_spair up_defined spair_defined
+
+lemmas sel_defined_iff_rules =
+  cfcomp2 sfst_defined_iff ssnd_defined_iff
+
 use "Tools/cont_consts.ML"
 use "Tools/cont_proc.ML"
 use "Tools/Domain/domain_constructors.ML"
--- a/src/HOLCF/Tools/Domain/domain_axioms.ML	Wed Feb 24 17:37:59 2010 -0800
+++ b/src/HOLCF/Tools/Domain/domain_axioms.ML	Thu Feb 25 13:16:28 2010 -0800
@@ -78,7 +78,7 @@
           (dc_abs oo (copy_of_dtyp map_tab r (dtyp_of_eq eqn)) oo dc_rep))
       end;
 
-(* -- definitions concerning the discriminators and selectors - *)
+(* -- definitions concerning the discriminators - *)
 
     val dis_defs = let
       fun ddef (con,_,_) = (dis_name con ^"_def",%%:(dis_name con) == 
@@ -115,13 +115,6 @@
           end
       in map pdef cons end;
 
-    val sel_defs = let
-      fun sdef con n arg = Option.map (fn sel => (sel^"_def",%%:sel == 
-                                                            list_ccomb(%%:(dname^"_when"),map 
-                                                                                            (fn (con', _, args) => if con'<>con then UU else
-                                                                                                               List.foldr /\# (Bound (length args - n)) args) cons))) (sel_of arg);
-    in map_filter I (maps (fn (con, _, args) => mapn (sdef con) 1 args) cons) end;
-
 
 (* ----- axiom and definitions concerning induction ------------------------- *)
 
@@ -141,7 +134,7 @@
   in (dnam,
       (if definitional then [] else [abs_iso_ax, rep_iso_ax, reach_ax]),
       (if definitional then [when_def] else [when_def, copy_def]) @
-      dis_defs @ mat_defs @ pat_defs @ sel_defs @
+      dis_defs @ mat_defs @ pat_defs @
       [take_def, finite_def])
   end; (* let (calc_axioms) *)
 
--- a/src/HOLCF/Tools/Domain/domain_constructors.ML	Wed Feb 24 17:37:59 2010 -0800
+++ b/src/HOLCF/Tools/Domain/domain_constructors.ML	Thu Feb 25 13:16:28 2010 -0800
@@ -14,7 +14,8 @@
       -> thm * thm
       -> theory
       -> { con_consts : term list,
-           con_defs : thm list }
+           con_defs : thm list,
+           sel_rews : thm list }
          * theory;
 end;
 
@@ -33,6 +34,7 @@
 fun mk_cfunT (T, U) = Type(@{type_name "->"}, [T, U]);
 
 infixr 6 ->>; val (op ->>) = mk_cfunT;
+infix -->>; val (op -->>) = Library.foldr mk_cfunT;
 
 fun dest_cfunT (Type(@{type_name "->"}, [T, U])) = (T, U)
   | dest_cfunT T = raise TYPE ("dest_cfunT", [T], []);
@@ -64,6 +66,25 @@
 
 infix 9 ` ; val (op `) = mk_capply;
 
+val list_ccomb : term * term list -> term = Library.foldl mk_capply;
+
+fun mk_ID T = Const (@{const_name ID}, T ->> T);
+
+fun cfcomp_const (T, U, V) =
+  Const (@{const_name cfcomp}, (U ->> V) ->> (T ->> U) ->> (T ->> V));
+
+fun mk_cfcomp (f, g) =
+  let
+    val (U, V) = dest_cfunT (Term.fastype_of f);
+    val (T, U') = dest_cfunT (Term.fastype_of g);
+  in
+    if U = U'
+    then mk_capply (mk_capply (cfcomp_const (T, U, V), f), g)
+    else raise TYPE ("mk_cfcomp", [U, U'], [f, g])
+  end;
+
+fun mk_bottom T = Const (@{const_name UU}, T);
+
 
 (*** Product type ***)
 
@@ -87,10 +108,18 @@
 
 fun mk_upT T = Type(@{type_name "u"}, [T]);
 
+fun dest_upT (Type(@{type_name "u"}, [T])) = T
+  | dest_upT T = raise TYPE ("dest_upT", [T], []);
+
 fun up_const T = Const(@{const_name up}, T ->> mk_upT T);
 
 fun mk_up t = up_const (Term.fastype_of t) ` t;
 
+fun fup_const (T, U) =
+  Const(@{const_name fup}, (T ->> U) ->> mk_upT T ->> U);
+
+fun from_up T = fup_const (T, T) ` mk_ID T;
+
 
 (*** Strict product type ***)
 
@@ -98,6 +127,9 @@
 
 fun mk_sprodT (T, U) = Type(@{type_name "**"}, [T, U]);
 
+fun dest_sprodT (Type(@{type_name "**"}, [T, U])) = (T, U)
+  | dest_sprodT T = raise TYPE ("dest_sprodT", [T], []);
+
 fun spair_const (T, U) =
   Const(@{const_name spair}, T ->> U ->> mk_sprodT (T, U));
 
@@ -110,6 +142,12 @@
   | mk_stuple (t::[]) = t
   | mk_stuple (t::ts) = mk_spair (t, mk_stuple ts);
 
+fun sfst_const (T, U) =
+  Const(@{const_name sfst}, mk_sprodT (T, U) ->> T);
+
+fun ssnd_const (T, U) =
+  Const(@{const_name ssnd}, mk_sprodT (T, U) ->> U);
+
 
 (*** Strict sum type ***)
 
@@ -139,6 +177,16 @@
     fst (inj (ts ~~ Ts))
   end;
 
+fun sscase_const (T, U, V) =
+  Const(@{const_name sscase},
+    (T ->> V) ->> (U ->> V) ->> mk_ssumT (T, U) ->> V);
+
+fun from_sinl (T, U) =
+  sscase_const (T, U, T) ` mk_ID T ` mk_bottom (U ->> T);
+
+fun from_sinr (T, U) =
+  sscase_const (T, U, U) ` mk_bottom (T ->> U) ` mk_ID U;
+
 
 (*** miscellaneous constructions ***)
 
@@ -151,7 +199,7 @@
     fun argTs (Type (_, Ts)) = Ts | argTs _ = [];
     fun auto T = T ->> T;
   in
-    Library.foldr mk_cfunT (map auto (argTs T), auto T)
+    map auto (argTs T) -->> auto T
   end;
 
 val mk_equals = Logic.mk_equals;
@@ -159,6 +207,13 @@
 val mk_trp = HOLogic.mk_Trueprop;
 val mk_fst = HOLogic.mk_fst;
 val mk_snd = HOLogic.mk_snd;
+val mk_not = HOLogic.mk_not;
+
+fun mk_strict t =
+  let val (T, U) = dest_cfunT (Term.fastype_of t);
+  in mk_eq (t ` mk_bottom T, mk_bottom U) end;
+
+fun mk_defined t = mk_not (mk_eq (t, mk_bottom (Term.fastype_of t)));
 
 fun mk_cont t =
   let val T = Term.fastype_of t
@@ -168,21 +223,6 @@
   let val (T, _) = dest_cfunT (Term.fastype_of t)
   in mk_capply (Const(@{const_name fix}, (T ->> T) ->> T), t) end;
 
-fun ID_const T = Const (@{const_name ID}, T ->> T);
-
-fun cfcomp_const (T, U, V) =
-  Const (@{const_name cfcomp}, (U ->> V) ->> (T ->> U) ->> (T ->> V));
-
-fun mk_cfcomp (f, g) =
-  let
-    val (U, V) = dest_cfunT (Term.fastype_of f);
-    val (T, U') = dest_cfunT (Term.fastype_of g);
-  in
-    if U = U'
-    then mk_capply (mk_capply (cfcomp_const (T, U, V), f), g)
-    else raise TYPE ("mk_cfcomp", [U, U'], [f, g])
-  end;
-
 fun mk_Rep_of T =
   Const (@{const_name Rep_of}, Term.itselfT T --> deflT) $ Logic.mk_type T;
 
@@ -196,8 +236,10 @@
 
 fun mk_eqs (t, u) = HOLogic.mk_Trueprop (HOLogic.mk_eq (t, u));
 
-(*** miscellaneous ***)
 
+(************************** miscellaneous functions ***************************)
+
+(*
 fun declare_consts
     (decls : (binding * typ * mixfix) list)
     (thy : theory)
@@ -208,6 +250,7 @@
   in
     (map con decls, thy)
   end;
+*)
 
 fun define_consts
     (specs : (binding * term * mixfix) list)
@@ -228,6 +271,16 @@
     ((consts, def_thms), thy)
   end;
 
+fun define_const
+    (spec : binding * term * mixfix)
+    (thy : theory)
+    : (term * thm) * theory =
+  let
+    val ((consts, def_thms), thy) = define_consts [spec] thy;
+  in
+    ((hd consts, hd def_thms), thy)
+  end;
+
 
 (************** generating beta reduction rules from definitions **************)
 
@@ -243,7 +296,7 @@
       let
         val (con, lam) = Logic.dest_equals (concl_of def_thm);
         val (args, rhs) = arglist lam;
-        val lhs = Library.foldl mk_capply (con, args);
+        val lhs = list_ccomb (con, args);
         val goal = mk_equals (lhs, rhs);
         val cs = ContProc.cont_thms lam;
         val betas = map (fn c => mk_meta_eq (c RS @{thm beta_cfun})) cs;
@@ -255,45 +308,211 @@
       end;
 end;
 
-(******************************* main function ********************************)
+(******************************************************************************)
+(************** definitions and theorems for selector functions ***************)
+(******************************************************************************)
+
+fun add_selectors
+    (spec : (term * (bool * binding option * typ) list) list)
+    (rep_const : term)
+    (abs_inv : thm)
+    (rep_strict : thm)
+    (rep_strict_iff : thm)
+    (con_betas : thm list)
+    (thy : theory)
+    : thm list * theory =
+  let
+
+    (* define selector functions *)
+    val ((sel_consts, sel_defs), thy) =
+      let
+        fun rangeT s = snd (dest_cfunT (Term.fastype_of s));
+        fun mk_outl s = mk_cfcomp (from_sinl (dest_ssumT (rangeT s)), s);
+        fun mk_outr s = mk_cfcomp (from_sinr (dest_ssumT (rangeT s)), s);
+        fun mk_sfst s = mk_cfcomp (sfst_const (dest_sprodT (rangeT s)), s);
+        fun mk_ssnd s = mk_cfcomp (ssnd_const (dest_sprodT (rangeT s)), s);
+        fun mk_down s = mk_cfcomp (from_up (dest_upT (rangeT s)), s);
+
+        fun sels_of_arg s (lazy, NONE,   T) = []
+          | sels_of_arg s (lazy, SOME b, T) =
+            [(b, if lazy then mk_down s else s, NoSyn)];
+        fun sels_of_args s [] = []
+          | sels_of_args s (v :: []) = sels_of_arg s v
+          | sels_of_args s (v :: vs) =
+            sels_of_arg (mk_sfst s) v @ sels_of_args (mk_ssnd s) vs;
+        fun sels_of_cons s [] = []
+          | sels_of_cons s ((con, args) :: []) = sels_of_args s args
+          | sels_of_cons s ((con, args) :: cs) =
+            sels_of_args (mk_outl s) args @ sels_of_cons (mk_outr s) cs;
+        val sel_eqns : (binding * term * mixfix) list =
+            sels_of_cons rep_const spec;
+      in
+        define_consts sel_eqns thy
+      end
+
+    (* replace bindings with terms in constructor spec *)
+    val spec2 : (term * (bool * term option * typ) list) list =
+      let
+        fun prep_arg (lazy, NONE, T) sels = ((lazy, NONE, T), sels)
+          | prep_arg (lazy, SOME _, T) sels =
+            ((lazy, SOME (hd sels), T), tl sels);
+        fun prep_con (con, args) sels =
+            apfst (pair con) (fold_map prep_arg args sels);
+      in
+        fst (fold_map prep_con spec sel_consts)
+      end;
+
+    (* prove selector strictness rules *)
+    val sel_stricts : thm list =
+      let
+        val rules = rep_strict :: sel_defs @ @{thms sel_strict_rules};
+        val tac = simp_tac (HOL_basic_ss addsimps rules) 1;
+        fun sel_strict sel =
+          let
+            val goal = mk_trp (mk_strict sel);
+          in
+            Goal.prove_global thy [] [] goal (K tac)
+          end
+      in
+        map sel_strict sel_consts
+      end
+
+    (* prove selector application rules *)
+    val sel_apps : thm list =
+      let
+        val rules = @{thms sel_app_rules};
+        val simps = simp_thms @ [abs_inv] @ con_betas @ sel_defs @ rules;
+        val tac = asm_simp_tac (HOL_basic_ss addsimps simps) 1;
+        fun sel_apps_of (i, (con, args)) =
+          let
+            val Ts : typ list = map #3 args;
+            val ns : string list = Datatype_Prop.make_tnames Ts;
+            val vs : term list = map Free (ns ~~ Ts);
+            val con_app : term = list_ccomb (con, vs);
+            val vs' : (bool * term) list = map #1 args ~~ vs;
+            fun one_same (n, sel, T) =
+              let
+                val xs = map snd (filter_out fst (nth_drop n vs'));
+                val assms = map (mk_trp o mk_defined) xs;
+                val concl = mk_trp (mk_eq (sel ` con_app, nth vs n));
+                val goal = Logic.list_implies (assms, concl);
+              in
+                Goal.prove_global thy [] [] goal (K tac)
+              end;
+            fun one_diff (n, sel, T) =
+              let
+                val goal = mk_trp (mk_eq (sel ` con_app, mk_bottom T));
+              in
+                Goal.prove_global thy [] [] goal (K tac)
+              end;
+            fun one_con (j, (_, args')) : thm list =
+              let
+                fun prep (i, (lazy, NONE, T)) = NONE
+                  | prep (i, (lazy, SOME sel, T)) = SOME (i, sel, T);
+                val sels : (int * term * typ) list =
+                  map_filter prep (map_index I args');
+              in
+                if i = j
+                then map one_same sels
+                else map one_diff sels
+              end
+          in
+            flat (map_index one_con spec2)
+          end
+      in
+        flat (map_index sel_apps_of spec2)
+      end
+
+  (* prove selector definedness rules *)
+    val sel_defins : thm list =
+      let
+        val rules = @{thms sel_defined_iff_rules};
+        val simps = rep_strict_iff :: sel_defs @ rules;
+        val tac = simp_tac (HOL_basic_ss addsimps simps) 1;
+        fun sel_defin sel =
+          let
+            val (T, U) = dest_cfunT (Term.fastype_of sel);
+            val x = Free ("x", T);
+            val lhs = mk_eq (sel ` x, mk_bottom U);
+            val rhs = mk_eq (x, mk_bottom T);
+            val goal = mk_trp (mk_eq (lhs, rhs));
+          in
+            Goal.prove_global thy [] [] goal (K tac)
+          end
+        fun one_arg (false, SOME sel, T) = SOME (sel_defin sel)
+          | one_arg _                    = NONE;
+      in
+        case spec2 of
+          [(con, args)] => map_filter one_arg args
+        | _             => []
+      end;
+
+  in
+    (sel_stricts @ sel_defins @ sel_apps, thy)
+  end
+
+(******************************************************************************)
+(************* definitions and theorems for constructor functions *************)
+(******************************************************************************)
 
 fun add_domain_constructors
     (dname : string)
     (lhsT : typ,
-     cons : (binding * (bool * binding option * typ) list * mixfix) list)
+     spec : (binding * (bool * binding option * typ) list * mixfix) list)
     (rep_const : term, abs_const : term)
     (rep_iso_thm : thm, abs_iso_thm : thm)
     (thy : theory) =
   let
-    (* TODO: re-enable this *)
-    (* val thy = Sign.add_path dname thy; *)
+
+    (* prove rep/abs strictness rules *)
+    val iso_locale = @{thm iso.intro} OF [abs_iso_thm, rep_iso_thm];
+    val rep_strict = iso_locale RS @{thm iso.rep_strict};
+    val abs_strict = iso_locale RS @{thm iso.abs_strict};
+    val rep_defined_iff = iso_locale RS @{thm iso.rep_defined_iff};
+    val abs_defined_iff = iso_locale RS @{thm iso.abs_defined_iff};
 
     (* define constructor functions *)
     val ((con_consts, con_def_thms), thy) =
       let
-        fun prep_con (bind, args, mx) =
-          (bind, args ~~ Datatype_Prop.make_tnames (map #3 args), mx);
-        fun var_of ((lazy, sel, T), name) = Free (name, T);
-        fun is_lazy ((lazy, sel, T), name) = lazy;
-        val cons' = map prep_con cons;
-        fun one_arg arg = (if is_lazy arg then mk_up else I) (var_of arg);
-        fun one_con (bind, args, mx) = mk_stuple (map one_arg args);
+        fun vars_of args =
+          let
+            val Ts = map (fn (lazy,sel,T) => T) args;
+            val ns = Datatype_Prop.make_tnames Ts;
+          in
+            map Free (ns ~~ Ts)
+          end;
+        fun one_arg (lazy,_,_) var = if lazy then mk_up var else var;
+        fun one_con (_,args,_) = mk_stuple (map2 one_arg args (vars_of args));
         fun mk_abs t = abs_const ` t;
-        val rhss = map mk_abs (mk_sinjects (map one_con cons'));
+        val rhss = map mk_abs (mk_sinjects (map one_con spec));
         fun mk_def (bind, args, mx) rhs =
-          (bind, big_lambdas (map var_of args) rhs, mx);
+          (bind, big_lambdas (vars_of args) rhs, mx);
       in
-        define_consts (map2 mk_def cons' rhss) thy
+        define_consts (map2 mk_def spec rhss) thy
       end;
 
+    (* prove beta reduction rules for constructors *)
     val con_beta_thms = map (beta_of_def thy) con_def_thms;
 
-    (* TODO: re-enable this *)
-    (* val thy = Sign.parent_path thy; *)
+    (* TODO: enable this earlier *)
+    val thy = Sign.add_path dname thy;
+
+    (* replace bindings with terms in constructor spec *)
+    val spec2 : (term * (bool * binding option * typ) list) list =
+      map2 (fn con => fn (b, args, mx) => (con, args)) con_consts spec;
+
+    (* define and prove theorems for selector functions *)
+    val (sel_thms : thm list, thy : theory) =
+      add_selectors spec2 rep_const
+        abs_iso_thm rep_strict rep_defined_iff con_beta_thms thy;
+
+    (* restore original signature path *)
+    val thy = Sign.parent_path thy;
 
     val result =
       { con_consts = con_consts,
-        con_defs = con_def_thms };
+        con_defs = con_def_thms,
+        sel_rews = sel_thms };
   in
     (result, thy)
   end;
--- a/src/HOLCF/Tools/Domain/domain_syntax.ML	Wed Feb 24 17:37:59 2010 -0800
+++ b/src/HOLCF/Tools/Domain/domain_syntax.ML	Thu Feb 25 13:16:28 2010 -0800
@@ -86,9 +86,6 @@
           (mat_name_ con,
            mk_matT(dtype, map third args, freetvar "t" 1),
            Mixfix(escape ("match_" ^ Binding.name_of con), [], Syntax.max_pri));
-      fun sel1 (_,sel,typ) =
-          Option.map (fn s => (s,dtype ->> typ,NoSyn)) sel;
-      fun sel (con,args,mx) = map_filter sel1 args;
       fun mk_patT (a,b)     = a ->> mk_maybeT b;
       fun pat_arg_typ n arg = mk_patT (third arg, freetvar "t" n);
       fun pat (con,args,mx) =
@@ -101,7 +98,6 @@
     val consts_dis = map dis cons';
     val consts_mat = map mat cons';
     val consts_pat = map pat cons';
-    val consts_sel = maps sel cons';
     end;
 
 (* ----- constants concerning induction ------------------------------------- *)
@@ -169,7 +165,7 @@
         if definitional then [] else [const_rep, const_abs, const_copy];
 
   in (optional_consts @ [const_when] @ 
-      consts_dis @ consts_mat @ consts_pat @ consts_sel @
+      consts_dis @ consts_mat @ consts_pat @
       [const_take, const_finite],
       (case_trans false :: case_trans true :: (abscon_trans false @ abscon_trans true @ Case_trans)))
   end; (* let *)
--- a/src/HOLCF/Tools/Domain/domain_theorems.ML	Wed Feb 24 17:37:59 2010 -0800
+++ b/src/HOLCF/Tools/Domain/domain_theorems.ML	Thu Feb 25 13:16:28 2010 -0800
@@ -161,6 +161,7 @@
   val axs_dis_def = map (get_def dis_name) cons;
   val axs_mat_def = map (get_def mat_name) cons;
   val axs_pat_def = map (get_def pat_name) cons;
+(*
   val axs_sel_def =
     let
       fun def_of_sel sel = ga (sel^"_def") dname;
@@ -169,6 +170,7 @@
     in
       maps defs_of_con cons
     end;
+*)
   val ax_copy_def = ga "copy_def" dname;
 end; (* local *)
 
@@ -196,6 +198,7 @@
     (rep_const, abs_const) (ax_rep_iso, ax_abs_iso) thy;
 
 val axs_con_def = #con_defs result;
+val sel_rews = #sel_rews result;
 
 (* ----- theorems concerning the isomorphism -------------------------------- *)
 
@@ -472,79 +475,6 @@
   val con_compacts = map con_compact cons;
 end;
 
-local
-  fun one_sel sel =
-    pg axs_sel_def (mk_trp (strict (%%:sel)))
-      (K [simp_tac (HOLCF_ss addsimps when_rews) 1]);
-
-  fun sel_strict (_, _, args) =
-    map_filter (Option.map one_sel o sel_of) args;
-in
-  val _ = trace " Proving sel_stricts...";
-  val sel_stricts = maps sel_strict cons;
-end;
-
-local
-  fun sel_app_same c n sel (con, args) =
-    let
-      val nlas = nonlazy args;
-      val vns = map vname args;
-      val vnn = List.nth (vns, n);
-      val nlas' = filter (fn v => v <> vnn) nlas;
-      val lhs = (%%:sel)`(con_app con args);
-      val goal = lift_defined %: (nlas', mk_trp (lhs === %:vnn));
-      fun tacs1 ctxt =
-        if vnn mem nlas
-                        (* FIXME! case_UU_tac *)
-        then [case_UU_tac ctxt (when_rews @ con_stricts) 1 vnn]
-        else [];
-      val tacs2 = [asm_simp_tac (HOLCF_ss addsimps when_rews) 1];
-    in pg axs_sel_def goal (fn ctxt => (tacs1 ctxt @ tacs2)) end;
-
-  fun sel_app_diff c n sel (con, args) =
-    let
-      val nlas = nonlazy args;
-      val goal = mk_trp (%%:sel ` con_app con args === UU);
-                        (* FIXME! case_UU_tac *)
-      fun tacs1 ctxt = map (case_UU_tac ctxt (when_rews @ con_stricts) 1) nlas;
-      val tacs2 = [asm_simp_tac (HOLCF_ss addsimps when_rews) 1];
-    in pg axs_sel_def goal (fn ctxt => (tacs1 ctxt @ tacs2)) end;
-
-  fun sel_app c n sel (con, _, args) =
-    if con = c
-    then sel_app_same c n sel (con, args)
-    else sel_app_diff c n sel (con, args);
-
-  fun one_sel c n sel = map (sel_app c n sel) cons;
-  fun one_sel' c n arg = Option.map (one_sel c n) (sel_of arg);
-  fun one_con (c, _, args) =
-    flat (map_filter I (mapn (one_sel' c) 0 args));
-in
-  val _ = trace " Proving sel_apps...";
-  val sel_apps = maps one_con cons;
-end;
-
-local
-  fun sel_defin sel =
-    let
-      val goal = defined (%:x_name) ==> defined (%%:sel`%x_name);
-      val tacs = [
-        rtac casedist 1,
-        contr_tac 1,
-        DETERM_UNTIL_SOLVED (CHANGED
-          (asm_simp_tac (HOLCF_ss addsimps sel_apps) 1))];
-    in pg [] goal (K tacs) end;
-in
-  val _ = trace " Proving sel_defins...";
-  val sel_defins =
-    if length cons = 1
-    then map_filter (fn arg => Option.map sel_defin (sel_of arg))
-                 (filter_out is_lazy (third (hd cons)))
-    else [];
-end;
-
-val sel_rews = sel_stricts @ sel_defins @ sel_apps;
-
 val _ = trace " Proving dist_les...";
 val dist_les =
   let