denesting of functions
authorhaftmann
Sun, 02 Jun 2013 10:57:21 +0200
changeset 52288 ca4932dad084
parent 52287 7e54c4d964e7
child 52289 83ce5d2841e7
denesting of functions
src/HOL/Tools/reification.ML
--- a/src/HOL/Tools/reification.ML	Sun Jun 02 09:10:53 2013 +0200
+++ b/src/HOL/Tools/reification.ML	Sun Jun 02 10:57:21 2013 +0200
@@ -103,179 +103,180 @@
 fun dereify ctxt eqs =
   rewrite_with ctxt (eqs @ @{thms nth_Cons_0 nth_Cons_Suc});
 
-fun conv ctxt eqs ct =
+fun index_of t bds =
+  let
+    val tt = HOLogic.listT (fastype_of t);
+  in
+    (case AList.lookup Type.could_unify bds tt of
+        NONE => error "index_of: type not found in environements!"
+      | SOME (tbs, tats) =>
+          let
+            val i = find_index (fn t' => t' = t) tats;
+            val j = find_index (fn t' => t' = t) tbs;
+          in
+            if j = ~1 then
+              if i = ~1
+              then (length tbs + length tats, AList.update Type.could_unify (tt, (tbs, tats @ [t])) bds)
+              else (i, bds)
+            else (j, bds)
+          end)
+  end;
+
+(* Generic decomp for reification : matches the actual term with the
+   rhs of one cong rule. The result of the matching guides the
+   proof synthesis: The matches of the introduced Variables A1 .. An are
+   processed recursively
+   The rest is instantiated in the cong rule,i.e. no reification is needed *)
+
+(* da is the decomposition for atoms, ie. it returns ([],g) where g
+   returns the right instance f (AtC n) = t , where AtC is the Atoms
+   constructor and n is the number of the atom corresponding to t *)
+fun decomp_reify da cgns (ct, ctxt) bds =
   let
-    fun index_of t bds =
+    val thy = Proof_Context.theory_of ctxt;
+    val cert = cterm_of thy;
+    val certT = ctyp_of thy;
+    fun tryabsdecomp (ct, ctxt) bds =
+      (case Thm.term_of ct of
+        Abs (_, xT, ta) =>
+          let
+            val ([raw_xn], ctxt') = Variable.variant_fixes ["x"] ctxt;
+            val (xn, ta) = Syntax_Trans.variant_abs (raw_xn, xT, ta);  (* FIXME !? *)
+            val x = Free (xn, xT);
+            val cx = cert x;
+            val cta = cert ta;
+            val bds = (case AList.lookup Type.could_unify bds (HOLogic.listT xT) of
+                NONE => error "tryabsdecomp: Type not found in the Environement"
+              | SOME (bsT, atsT) => AList.update Type.could_unify (HOLogic.listT xT,
+                  (x :: bsT, atsT)) bds);
+           in (([(cta, ctxt')],
+                fn ([th], bds) =>
+                  (hd (Variable.export ctxt' ctxt [(Thm.forall_intr cx th) COMP allI]),
+                   let
+                     val (bsT, asT) = the (AList.lookup Type.could_unify bds (HOLogic.listT xT));
+                   in
+                     AList.update Type.could_unify (HOLogic.listT xT, (tl bsT, asT)) bds
+                   end)),
+               bds)
+           end
+       | _ => da (ct, ctxt) bds)
+  in
+    (case cgns of
+      [] => tryabsdecomp (ct, ctxt) bds
+    | ((vns, cong) :: congs) =>
+        (let
+          val (tyenv, tmenv) =
+            Pattern.match thy
+              ((fst o HOLogic.dest_eq o HOLogic.dest_Trueprop) (concl_of cong), Thm.term_of ct)
+              (Vartab.empty, Vartab.empty);
+          val (fnvs, invs) = List.partition (fn ((vn, _),_) => member (op =) vns vn) (Vartab.dest tmenv);
+          val (fts, its) =
+            (map (snd o snd) fnvs,
+             map (fn ((vn, vi), (tT, t)) => (cert (Var ((vn, vi), tT)), cert t)) invs);
+          val ctyenv = map (fn ((vn, vi), (s, ty)) => (certT (TVar((vn, vi), s)), certT ty)) (Vartab.dest tyenv);
+        in
+          ((map cert fts ~~ replicate (length fts) ctxt,
+             apfst (FWD (Drule.instantiate_normalize (ctyenv, its) cong))), bds)
+        end handle Pattern.MATCH => decomp_reify da congs (ct, ctxt) bds))
+  end;
+
+fun get_nths (t as (Const (@{const_name "List.nth"}, _) $ vs $ n)) =
+      AList.update (op aconv) (t, (vs, n))
+  | get_nths (t1 $ t2) = get_nths t1 #> get_nths t2
+  | get_nths (Abs (_, _, t')) = get_nths t'
+  | get_nths _ = I;
+
+fun tryeqs [] (ct, ctxt) bds = error "Cannot find the atoms equation"
+  | tryeqs (eq :: eqs) (ct, ctxt) bds = ((
       let
-        val tt = HOLogic.listT (fastype_of t);
-      in
-        (case AList.lookup Type.could_unify bds tt of
-            NONE => error "index_of: type not found in environements!"
-          | SOME (tbs, tats) =>
-              let
-                val i = find_index (fn t' => t' = t) tats;
-                val j = find_index (fn t' => t' = t) tbs;
-              in
-                if j = ~1 then
-                  if i = ~1
-                  then (length tbs + length tats, AList.update Type.could_unify (tt, (tbs, tats @ [t])) bds)
-                  else (i, bds)
-                else (j, bds)
-              end)
-      end;
-
-    (* Generic decomp for reification : matches the actual term with the
-       rhs of one cong rule. The result of the matching guides the
-       proof synthesis: The matches of the introduced Variables A1 .. An are
-       processed recursively
-       The rest is instantiated in the cong rule,i.e. no reification is needed *)
-
-    (* da is the decomposition for atoms, ie. it returns ([],g) where g
-       returns the right instance f (AtC n) = t , where AtC is the Atoms
-       constructor and n is the number of the atom corresponding to t *)
-    fun decomp_reify da cgns (ct, ctxt) bds =
-      let
-        val thy = Proof_Context.theory_of ctxt;
+        val rhs = eq |> prop_of |> HOLogic.dest_Trueprop  |> HOLogic.dest_eq |> snd;
+        val nths = get_nths rhs [];
+        val (vss, _) = fold_rev (fn (_, (vs, n)) => fn (vss, ns) =>
+          (insert (op aconv) vs vss, insert (op aconv) n ns)) nths ([], []);
+        val (vsns, ctxt') = Variable.variant_fixes (replicate (length vss) "vs") ctxt;
+        val (xns, ctxt'') = Variable.variant_fixes (replicate (length nths) "x") ctxt';
+        val thy = Proof_Context.theory_of ctxt'';
         val cert = cterm_of thy;
         val certT = ctyp_of thy;
-        fun tryabsdecomp (ct, ctxt) bds =
-          (case Thm.term_of ct of
-            Abs (_, xT, ta) =>
-              let
-                val ([raw_xn], ctxt') = Variable.variant_fixes ["x"] ctxt;
-                val (xn, ta) = Syntax_Trans.variant_abs (raw_xn, xT, ta);  (* FIXME !? *)
-                val x = Free (xn, xT);
-                val cx = cert x;
-                val cta = cert ta;
-                val bds = (case AList.lookup Type.could_unify bds (HOLogic.listT xT) of
-                    NONE => error "tryabsdecomp: Type not found in the Environement"
-                  | SOME (bsT, atsT) => AList.update Type.could_unify (HOLogic.listT xT,
-                      (x :: bsT, atsT)) bds);
-               in (([(cta, ctxt')],
-                    fn ([th], bds) =>
-                      (hd (Variable.export ctxt' ctxt [(Thm.forall_intr cx th) COMP allI]),
-                       let
-                         val (bsT, asT) = the (AList.lookup Type.could_unify bds (HOLogic.listT xT));
-                       in
-                         AList.update Type.could_unify (HOLogic.listT xT, (tl bsT, asT)) bds
-                       end)),
-                   bds)
-               end
-           | _ => da (ct, ctxt) bds)
-      in
-        (case cgns of
-          [] => tryabsdecomp (ct, ctxt) bds
-        | ((vns, cong) :: congs) =>
-            (let
-              val (tyenv, tmenv) =
-                Pattern.match thy
-                  ((fst o HOLogic.dest_eq o HOLogic.dest_Trueprop) (concl_of cong), Thm.term_of ct)
-                  (Vartab.empty, Vartab.empty);
-              val (fnvs, invs) = List.partition (fn ((vn, _),_) => member (op =) vns vn) (Vartab.dest tmenv);
-              val (fts, its) =
-                (map (snd o snd) fnvs,
-                 map (fn ((vn, vi), (tT, t)) => (cert (Var ((vn, vi), tT)), cert t)) invs);
-              val ctyenv = map (fn ((vn, vi), (s, ty)) => (certT (TVar((vn, vi), s)), certT ty)) (Vartab.dest tyenv);
-            in
-              ((map cert fts ~~ replicate (length fts) ctxt,
-                 apfst (FWD (Drule.instantiate_normalize (ctyenv, its) cong))), bds)
-            end handle Pattern.MATCH => decomp_reify da congs (ct, ctxt) bds))
-      end;
-
- (* looks for the atoms equation and instantiates it with the right number *)
-    fun mk_decompatom eqs (ct, ctxt) bds = (([], fn (_, bds) =>
-      let
-        val tT = fastype_of (Thm.term_of ct);
-        fun isat eq =
+        val vsns_map = vss ~~ vsns;
+        val xns_map = fst (split_list nths) ~~ xns;
+        val subst = map (fn (nt, xn) => (nt, Var ((xn, 0), fastype_of nt))) xns_map;
+        val rhs_P = subst_free subst rhs;
+        val (tyenv, tmenv) = Pattern.match thy (rhs_P, Thm.term_of ct) (Vartab.empty, Vartab.empty);
+        val sbst = Envir.subst_term (tyenv, tmenv);
+        val sbsT = Envir.subst_type tyenv;
+        val subst_ty = map (fn (n, (s, t)) =>
+          (certT (TVar (n, s)), certT t)) (Vartab.dest tyenv)
+        val tml = Vartab.dest tmenv;
+        val (subst_ns, bds) = fold_map
+          (fn (Const _ $ _ $ n, Var (xn0, _)) => fn bds =>
+            let
+              val name = snd (the (AList.lookup (op =) tml xn0));
+              val (idx, bds) = index_of name bds;
+            in ((cert n, idx |> (HOLogic.mk_nat #> cert)), bds) end) subst bds;
+        val subst_vs =
           let
-            val rhs = eq |> prop_of |> HOLogic.dest_Trueprop |> HOLogic.dest_eq |> snd;
-          in exists_Const
-            (fn (n, ty) => n = @{const_name "List.nth"}
-              andalso AList.defined Type.could_unify bds (domain_type ty)) rhs
-              andalso Type.could_unify (fastype_of rhs, tT)
-          end;
-
-        fun get_nths (t as (Const (@{const_name "List.nth"}, _) $ vs $ n)) =
-              AList.update (op aconv) (t, (vs, n))
-          | get_nths (t1 $ t2) = get_nths t1 #> get_nths t2
-          | get_nths (Abs (_, _, t')) = get_nths t'
-          | get_nths _ = I;
-
-        fun tryeqs [] bds = error "Cannot find the atoms equation"
-          | tryeqs (eq :: eqs) bds = ((
+            fun h (Const _ $ (vs as Var (_, lT)) $ _, Var (_, T)) =
               let
-                val rhs = eq |> prop_of |> HOLogic.dest_Trueprop  |> HOLogic.dest_eq |> snd;
-                val nths = get_nths rhs [];
-                val (vss, _) = fold_rev (fn (_, (vs, n)) => fn (vss, ns) =>
-                  (insert (op aconv) vs vss, insert (op aconv) n ns)) nths ([], []);
-                val (vsns, ctxt') = Variable.variant_fixes (replicate (length vss) "vs") ctxt;
-                val (xns, ctxt'') = Variable.variant_fixes (replicate (length nths) "x") ctxt';
-                val thy = Proof_Context.theory_of ctxt'';
-                val cert = cterm_of thy;
-                val certT = ctyp_of thy;
-                val vsns_map = vss ~~ vsns;
-                val xns_map = fst (split_list nths) ~~ xns;
-                val subst = map (fn (nt, xn) => (nt, Var ((xn, 0), fastype_of nt))) xns_map;
-                val rhs_P = subst_free subst rhs;
-                val (tyenv, tmenv) = Pattern.match thy (rhs_P, Thm.term_of ct) (Vartab.empty, Vartab.empty);
-                val sbst = Envir.subst_term (tyenv, tmenv);
-                val sbsT = Envir.subst_type tyenv;
-                val subst_ty = map (fn (n, (s, t)) =>
-                  (certT (TVar (n, s)), certT t)) (Vartab.dest tyenv)
-                val tml = Vartab.dest tmenv;
-                val (subst_ns, bds) = fold_map
-                  (fn (Const _ $ _ $ n, Var (xn0, _)) => fn bds =>
-                    let
-                      val name = snd (the (AList.lookup (op =) tml xn0));
-                      val (idx, bds) = index_of name bds;
-                    in ((cert n, idx |> (HOLogic.mk_nat #> cert)), bds) end) subst bds;
-                val subst_vs =
-                  let
-                    fun h (Const _ $ (vs as Var (_, lT)) $ _, Var (_, T)) =
-                      let
-                        val cns = sbst (Const (@{const_name "List.Cons"}, T --> lT --> lT));
-                        val lT' = sbsT lT;
-                        val (bsT, _) = the (AList.lookup Type.could_unify bds lT);
-                        val vsn = the (AList.lookup (op =) vsns_map vs);
-                        val cvs = cert (fold_rev (fn x => fn xs => cns $ x $xs) bsT (Free (vsn, lT')));
-                      in (cert vs, cvs) end;
-                  in map h subst end;
-                val cts = map (fn ((vn, vi), (tT, t)) => (cert (Var ((vn, vi), tT)), cert t))
-                  (fold (AList.delete (fn (((a : string), _), (b, _)) => a = b))
-                    (map (fn n => (n, 0)) xns) tml);
-                val substt =
-                  let
-                    val ih = Drule.cterm_rule (Thm.instantiate (subst_ty, []));
-                  in map (pairself ih) (subst_ns @ subst_vs @ cts) end;
-                val th = (Drule.instantiate_normalize (subst_ty, substt) eq) RS sym;
-              in (hd (Variable.export ctxt'' ctxt [th]), bds) end)
-              handle Pattern.MATCH => tryeqs eqs bds)
-          in tryeqs (filter isat eqs) bds end), bds);
+                val cns = sbst (Const (@{const_name "List.Cons"}, T --> lT --> lT));
+                val lT' = sbsT lT;
+                val (bsT, _) = the (AList.lookup Type.could_unify bds lT);
+                val vsn = the (AList.lookup (op =) vsns_map vs);
+                val cvs = cert (fold_rev (fn x => fn xs => cns $ x $xs) bsT (Free (vsn, lT')));
+              in (cert vs, cvs) end;
+          in map h subst end;
+        val cts = map (fn ((vn, vi), (tT, t)) => (cert (Var ((vn, vi), tT)), cert t))
+          (fold (AList.delete (fn (((a : string), _), (b, _)) => a = b))
+            (map (fn n => (n, 0)) xns) tml);
+        val substt =
+          let
+            val ih = Drule.cterm_rule (Thm.instantiate (subst_ty, []));
+          in map (pairself ih) (subst_ns @ subst_vs @ cts) end;
+        val th = (Drule.instantiate_normalize (subst_ty, substt) eq) RS sym;
+      in (hd (Variable.export ctxt'' ctxt [th]), bds) end)
+      handle Pattern.MATCH => tryeqs eqs (ct, ctxt) bds);
+
+(* looks for the atoms equation and instantiates it with the right number *)
 
-  (* Generic reification procedure: *)
-  (* creates all needed cong rules and then just uses the theorem synthesis *)
-
-    fun mk_congs ctxt eqs =
+fun mk_decompatom eqs (ct, ctxt) bds = (([], fn (_, bds) =>
+  let
+    val tT = fastype_of (Thm.term_of ct);
+    fun isat eq =
       let
-        val fs = fold_rev (fn eq => insert (op =) (eq |> prop_of |> HOLogic.dest_Trueprop
-          |> HOLogic.dest_eq |> fst |> strip_comb
-          |> fst)) eqs [];
-        val tys = fold_rev (fn f => fold (insert (op =)) (f |> fastype_of |> binder_types |> tl)) fs [];
-        val (vs, ctxt') = Variable.variant_fixes (replicate (length tys) "vs") ctxt;
-        val cert = cterm_of (Proof_Context.theory_of ctxt');
-        val subst =
-          the o AList.lookup (op =) (map2 (fn T => fn v => (T, cert (Free (v, T)))) tys vs);
-        fun prep_eq eq =
-          let
-            val (_, _ :: vs) = eq |> prop_of |> HOLogic.dest_Trueprop
-              |> HOLogic.dest_eq |> fst |> strip_comb;
-            val subst = map_filter (fn (v as Var (_, T)) => SOME (cert v, subst T)
-              | _ => NONE) vs;
-          in Thm.instantiate ([], subst) eq end;
-        val (ps, congs) = map_split (mk_congeq ctxt' fs o prep_eq) eqs;
-        val bds = AList.make (K ([], [])) tys;
-      in (ps ~~ Variable.export ctxt' ctxt congs, bds) end
+        val rhs = eq |> prop_of |> HOLogic.dest_Trueprop |> HOLogic.dest_eq |> snd;
+      in exists_Const
+        (fn (n, ty) => n = @{const_name "List.nth"}
+          andalso AList.defined Type.could_unify bds (domain_type ty)) rhs
+          andalso Type.could_unify (fastype_of rhs, tT)
+      end;
+  in tryeqs (filter isat eqs) (ct, ctxt) bds end), bds);
+
+(* Generic reification procedure: *)
+(* creates all needed cong rules and then just uses the theorem synthesis *)
 
+fun mk_congs ctxt eqs =
+  let
+    val fs = fold_rev (fn eq => insert (op =) (eq |> prop_of |> HOLogic.dest_Trueprop
+      |> HOLogic.dest_eq |> fst |> strip_comb
+      |> fst)) eqs [];
+    val tys = fold_rev (fn f => fold (insert (op =)) (f |> fastype_of |> binder_types |> tl)) fs [];
+    val (vs, ctxt') = Variable.variant_fixes (replicate (length tys) "vs") ctxt;
+    val cert = cterm_of (Proof_Context.theory_of ctxt');
+    val subst =
+      the o AList.lookup (op =) (map2 (fn T => fn v => (T, cert (Free (v, T)))) tys vs);
+    fun prep_eq eq =
+      let
+        val (_, _ :: vs) = eq |> prop_of |> HOLogic.dest_Trueprop
+          |> HOLogic.dest_eq |> fst |> strip_comb;
+        val subst = map_filter (fn (v as Var (_, T)) => SOME (cert v, subst T)
+          | _ => NONE) vs;
+      in Thm.instantiate ([], subst) eq end;
+    val (ps, congs) = map_split (mk_congeq ctxt' fs o prep_eq) eqs;
+    val bds = AList.make (K ([], [])) tys;
+  in (ps ~~ Variable.export ctxt' ctxt congs, bds) end
+
+fun conv ctxt eqs ct =
+  let
     val (congs, bds) = mk_congs ctxt eqs;
     val congs = rearrange congs;
     val (th, bds') = apfst mk_eq (divide_and_conquer' (decomp_reify (mk_decompatom eqs) congs) (ct, ctxt) bds);