src/HOLCF/Tools/fixrec_package.ML
author wenzelm
Thu, 30 Aug 2007 22:35:34 +0200
changeset 24493 d4380e9b287b
parent 23779 742be2833ccd
child 24680 0d355aa59e67
permissions -rw-r--r--
replaced ProofContext.infer_types by general Syntax.check_terms;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
     1
(*  Title:      HOLCF/Tools/fixrec_package.ML
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
     2
    ID:         $Id$
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
     3
    Author:     Amber Telfer and Brian Huffman
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
     4
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
     5
Recursive function definition package for HOLCF.
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
     6
*)
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
     7
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
     8
signature FIXREC_PACKAGE =
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
     9
sig
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    10
  val legacy_infer_term: theory -> term -> term
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    11
  val legacy_infer_prop: theory -> term -> term
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    12
  val add_fixrec: bool -> ((string * Attrib.src list) * string) list list -> theory -> theory
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    13
  val add_fixrec_i: bool -> ((string * attribute list) * term) list list -> theory -> theory
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    14
  val add_fixpat: (string * Attrib.src list) * string list -> theory -> theory
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    15
  val add_fixpat_i: (string * attribute list) * term list -> theory -> theory
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    16
end;
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    17
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    18
structure FixrecPackage: FIXREC_PACKAGE =
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    19
struct
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    20
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    21
(* legacy type inference *)
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    22
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    23
fun legacy_infer_term thy t =
24493
d4380e9b287b replaced ProofContext.infer_types by general Syntax.check_terms;
wenzelm
parents: 23779
diff changeset
    24
  singleton (Syntax.check_terms (ProofContext.init thy)) (Sign.intern_term thy t);
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    25
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    26
fun legacy_infer_prop thy t = legacy_infer_term thy (TypeInfer.constrain t propT);
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    27
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    28
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    29
val fix_eq2 = thm "fix_eq2";
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    30
val def_fix_ind = thm "def_fix_ind";
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    31
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    32
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    33
fun fixrec_err s = error ("fixrec definition error:\n" ^ s);
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    34
fun fixrec_eq_err thy s eq =
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    35
  fixrec_err (s ^ "\nin\n" ^ quote (Sign.string_of_term thy eq));
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    36
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    37
(* ->> is taken from holcf_logic.ML *)
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    38
(* TODO: fix dependencies so we can import HOLCFLogic here *)
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    39
infixr 6 ->>;
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    40
fun S ->> T = Type ("Cfun.->",[S,T]);
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    41
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    42
(* extern_name is taken from domain/library.ML *)
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    43
fun extern_name con = case Symbol.explode con of 
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    44
		   ("o"::"p"::" "::rest) => implode rest
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    45
		   | _ => con;
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    46
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    47
val mk_trp = HOLogic.mk_Trueprop;
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    48
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    49
(* splits a cterm into the right and lefthand sides of equality *)
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    50
fun dest_eqs t = HOLogic.dest_eq (HOLogic.dest_Trueprop t);
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    51
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    52
(* similar to Thm.head_of, but for continuous application *)
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    53
fun chead_of (Const("Cfun.Rep_CFun",_)$f$t) = chead_of f
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    54
  | chead_of u = u;
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    55
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    56
(* these are helpful functions copied from HOLCF/domain/library.ML *)
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    57
fun %: s = Free(s,dummyT);
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    58
fun %%: s = Const(s,dummyT);
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    59
infix 0 ==;  fun S ==  T = %%:"==" $ S $ T;
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    60
infix 1 ===; fun S === T = %%:"op =" $ S $ T;
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    61
infix 9 `  ; fun f ` x = %%:"Cfun.Rep_CFun" $ f $ x;
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    62
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    63
(* builds the expression (LAM v. rhs) *)
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    64
fun big_lambda v rhs = %%:"Cfun.Abs_CFun"$(Term.lambda v rhs);
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    65
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    66
(* builds the expression (LAM v1 v2 .. vn. rhs) *)
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    67
fun big_lambdas [] rhs = rhs
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    68
  | big_lambdas (v::vs) rhs = big_lambda v (big_lambdas vs rhs);
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    69
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    70
(* builds the expression (LAM <v1,v2,..,vn>. rhs) *)
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    71
fun lambda_ctuple [] rhs = big_lambda (%:"unit") rhs
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    72
  | lambda_ctuple (v::[]) rhs = big_lambda v rhs
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    73
  | lambda_ctuple (v::vs) rhs =
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    74
      %%:"Cprod.csplit"`(big_lambda v (lambda_ctuple vs rhs));
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    75
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    76
(* builds the expression <v1,v2,..,vn> *)
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    77
fun mk_ctuple [] = %%:"UU"
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    78
|   mk_ctuple (t::[]) = t
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    79
|   mk_ctuple (t::ts) = %%:"Cprod.cpair"`t`(mk_ctuple ts);
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    80
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    81
(*************************************************************************)
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    82
(************* fixed-point definitions and unfolding theorems ************)
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    83
(*************************************************************************)
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    84
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    85
fun add_fixdefs eqs thy =
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    86
  let
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    87
    val (lhss,rhss) = ListPair.unzip (map dest_eqs eqs);
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    88
    val fixpoint = %%:"Fix.fix"`lambda_ctuple lhss (mk_ctuple rhss);
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    89
    
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    90
    fun one_def (l as Const(n,T)) r =
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    91
          let val b = Sign.base_name n in (b, (b^"_def", l == r)) end
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    92
      | one_def _ _ = fixrec_err "fixdefs: lhs not of correct form";
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    93
    fun defs [] _ = []
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    94
      | defs (l::[]) r = [one_def l r]
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    95
      | defs (l::ls) r = one_def l (%%:"Cprod.cfst"`r) :: defs ls (%%:"Cprod.csnd"`r);
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    96
    val (names, pre_fixdefs) = ListPair.unzip (defs lhss fixpoint);
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    97
    
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    98
    val fixdefs = map (apsnd (legacy_infer_prop thy)) pre_fixdefs;
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    99
    val (fixdef_thms, thy') =
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   100
      PureThy.add_defs_i false (map Thm.no_attributes fixdefs) thy;
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   101
    val ctuple_fixdef_thm = foldr1 (fn (x,y) => cpair_equalI OF [x,y]) fixdef_thms;
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   102
    
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   103
    val ctuple_unfold = legacy_infer_term thy' (mk_trp (mk_ctuple lhss === mk_ctuple rhss));
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   104
    val ctuple_unfold_thm = Goal.prove_global thy' [] [] ctuple_unfold
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   105
          (fn _ => EVERY [rtac (ctuple_fixdef_thm RS fix_eq2 RS trans) 1,
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   106
                    simp_tac (simpset_of thy') 1]);
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   107
    val ctuple_induct_thm =
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   108
          (space_implode "_" names ^ "_induct", ctuple_fixdef_thm RS def_fix_ind);
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   109
    
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   110
    fun unfolds [] thm = []
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   111
      | unfolds (n::[]) thm = [(n^"_unfold", thm)]
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   112
      | unfolds (n::ns) thm = let
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   113
          val thmL = thm RS cpair_eqD1;
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   114
          val thmR = thm RS cpair_eqD2;
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   115
        in (n^"_unfold", thmL) :: unfolds ns thmR end;
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   116
    val unfold_thms = unfolds names ctuple_unfold_thm;
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   117
    val thms = ctuple_induct_thm :: unfold_thms;
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   118
    val (_, thy'') = PureThy.add_thms (map Thm.no_attributes thms) thy';
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   119
  in
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   120
    (thy'', names, fixdef_thms, map snd unfold_thms)
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   121
  end;
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   122
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   123
(*************************************************************************)
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   124
(*********** monadic notation and pattern matching compilation ***********)
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   125
(*************************************************************************)
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   126
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   127
fun add_names (Const(a,_), bs) = insert (op =) (Sign.base_name a) bs
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   128
  | add_names (Free(a,_) , bs) = insert (op =) a bs
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   129
  | add_names (f $ u     , bs) = add_names (f, add_names(u, bs))
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   130
  | add_names (Abs(a,_,t), bs) = add_names (t, insert (op =) a bs)
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   131
  | add_names (_         , bs) = bs;
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   132
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   133
fun add_terms ts xs = foldr add_names xs ts;
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   134
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   135
(* builds a monadic term for matching a constructor pattern *)
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   136
fun pre_build pat rhs vs taken =
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   137
  case pat of
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   138
    Const("Cfun.Rep_CFun",_)$f$(v as Free(n,T)) =>
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   139
      pre_build f rhs (v::vs) taken
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   140
  | Const("Cfun.Rep_CFun",_)$f$x =>
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   141
      let val (rhs', v, taken') = pre_build x rhs [] taken;
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   142
      in pre_build f rhs' (v::vs) taken' end
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   143
  | Const(c,T) =>
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   144
      let
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   145
        val n = Name.variant taken "v";
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   146
        fun result_type (Type("Cfun.->",[_,T])) (x::xs) = result_type T xs
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   147
          | result_type T _ = T;
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   148
        val v = Free(n, result_type T vs);
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   149
        val m = "match_"^(extern_name(Sign.base_name c));
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   150
        val k = lambda_ctuple vs rhs;
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   151
      in
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   152
        (%%:"Fixrec.bind"`(%%:m`v)`k, v, n::taken)
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   153
      end
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   154
  | Free(n,_) => fixrec_err ("expected constructor, found free variable " ^ quote n)
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   155
  | _ => fixrec_err "pre_build: invalid pattern";
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   156
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   157
(* builds a monadic term for matching a function definition pattern *)
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   158
(* returns (name, arity, matcher) *)
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   159
fun building pat rhs vs taken =
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   160
  case pat of
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   161
    Const("Cfun.Rep_CFun", _)$f$(v as Free(n,T)) =>
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   162
      building f rhs (v::vs) taken
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   163
  | Const("Cfun.Rep_CFun", _)$f$x =>
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   164
      let val (rhs', v, taken') = pre_build x rhs [] taken;
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   165
      in building f rhs' (v::vs) taken' end
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   166
  | Const(name,_) => (name, length vs, big_lambdas vs rhs)
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   167
  | _ => fixrec_err "function is not declared as constant in theory";
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   168
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   169
fun match_eq eq = 
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   170
  let val (lhs,rhs) = dest_eqs eq;
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   171
  in building lhs (%%:"Fixrec.return"`rhs) [] (add_terms [eq] []) end;
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   172
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   173
(* returns the sum (using +++) of the terms in ms *)
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   174
(* also applies "run" to the result! *)
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   175
fun fatbar arity ms =
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   176
  let
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   177
    fun unLAM 0 t = t
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   178
      | unLAM n (_$Abs(_,_,t)) = unLAM (n-1) t
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   179
      | unLAM _ _ = fixrec_err "fatbar: internal error, not enough LAMs";
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   180
    fun reLAM 0 t = t
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   181
      | reLAM n t = reLAM (n-1) (%%:"Cfun.Abs_CFun" $ Abs("",dummyT,t));
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   182
    fun mplus (x,y) = %%:"Fixrec.mplus"`x`y;
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   183
    val msum = foldr1 mplus (map (unLAM arity) ms);
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   184
  in
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   185
    reLAM arity (%%:"Fixrec.run"`msum)
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   186
  end;
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   187
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   188
fun unzip3 [] = ([],[],[])
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   189
  | unzip3 ((x,y,z)::ts) =
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   190
      let val (xs,ys,zs) = unzip3 ts
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   191
      in (x::xs, y::ys, z::zs) end;
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   192
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   193
(* this is the pattern-matching compiler function *)
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   194
fun compile_pats eqs = 
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   195
  let
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   196
    val ((n::names),(a::arities),mats) = unzip3 (map match_eq eqs);
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   197
    val cname = if forall (fn x => n=x) names then n
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   198
          else fixrec_err "all equations in block must define the same function";
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   199
    val arity = if forall (fn x => a=x) arities then a
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   200
          else fixrec_err "all equations in block must have the same arity";
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   201
    val rhs = fatbar arity mats;
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   202
  in
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   203
    mk_trp (%%:cname === rhs)
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   204
  end;
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   205
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   206
(*************************************************************************)
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   207
(********************** Proving associated theorems **********************)
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   208
(*************************************************************************)
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   209
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   210
(* proves a block of pattern matching equations as theorems, using unfold *)
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   211
fun make_simps thy (unfold_thm, eqns) =
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   212
  let
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   213
    val tacs = [rtac (unfold_thm RS ssubst_lhs) 1, asm_simp_tac (simpset_of thy) 1];
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   214
    fun prove_term t = Goal.prove_global thy [] [] t (K (EVERY tacs));
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   215
    fun prove_eqn ((name, eqn_t), atts) = ((name, prove_term eqn_t), atts);
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   216
  in
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   217
    map prove_eqn eqns
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   218
  end;
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   219
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   220
(*************************************************************************)
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   221
(************************* Main fixrec function **************************)
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   222
(*************************************************************************)
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   223
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   224
fun gen_add_fixrec prep_prop prep_attrib strict blocks thy =
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   225
  let
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   226
    val eqns = List.concat blocks;
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   227
    val lengths = map length blocks;
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   228
    
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   229
    val ((names, srcss), strings) = apfst split_list (split_list eqns);
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   230
    val atts = map (map (prep_attrib thy)) srcss;
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   231
    val eqn_ts = map (prep_prop thy) strings;
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   232
    val rec_ts = map (fn eq => chead_of (fst (dest_eqs (Logic.strip_imp_concl eq)))
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   233
      handle TERM _ => fixrec_eq_err thy "not a proper equation" eq) eqn_ts;
23779
742be2833ccd Function unify_consts moved from OldInductivePackage to PrimrecPackage.
berghofe
parents: 23152
diff changeset
   234
    val (_, eqn_ts') = PrimrecPackage.unify_consts thy rec_ts eqn_ts;
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   235
    
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   236
    fun unconcat [] _ = []
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   237
      | unconcat (n::ns) xs = List.take (xs,n) :: unconcat ns (List.drop (xs,n));
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   238
    val pattern_blocks = unconcat lengths (map Logic.strip_imp_concl eqn_ts');
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   239
    val compiled_ts = map (legacy_infer_term thy o compile_pats) pattern_blocks;
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   240
    val (thy', cnames, fixdef_thms, unfold_thms) = add_fixdefs compiled_ts thy;
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   241
  in
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   242
    if strict then let (* only prove simp rules if strict = true *)
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   243
      val eqn_blocks = unconcat lengths ((names ~~ eqn_ts') ~~ atts);
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   244
      val simps = List.concat (map (make_simps thy') (unfold_thms ~~ eqn_blocks));
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   245
      val (simp_thms, thy'') = PureThy.add_thms simps thy';
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   246
      
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   247
      val simp_names = map (fn name => name^"_simps") cnames;
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   248
      val simp_attribute = rpair [Simplifier.simp_add];
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   249
      val simps' = map simp_attribute (simp_names ~~ unconcat lengths simp_thms);
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   250
    in
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   251
      (snd o PureThy.add_thmss simps') thy''
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   252
    end
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   253
    else thy'
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   254
  end;
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   255
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   256
val add_fixrec = gen_add_fixrec Sign.read_prop Attrib.attribute;
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   257
val add_fixrec_i = gen_add_fixrec Sign.cert_prop (K I);
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   258
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   259
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   260
(*************************************************************************)
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   261
(******************************** Fixpat *********************************)
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   262
(*************************************************************************)
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   263
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   264
fun fix_pat thy t = 
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   265
  let
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   266
    val T = fastype_of t;
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   267
    val eq = mk_trp (HOLogic.eq_const T $ t $ Var (("x",0),T));
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   268
    val cname = case chead_of t of Const(c,_) => c | _ =>
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   269
              fixrec_err "function is not declared as constant in theory";
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   270
    val unfold_thm = PureThy.get_thm thy (Name (cname^"_unfold"));
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   271
    val simp = Goal.prove_global thy [] [] eq
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   272
          (fn _ => EVERY [stac unfold_thm 1, simp_tac (simpset_of thy) 1]);
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   273
  in simp end;
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   274
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   275
fun gen_add_fixpat prep_term prep_attrib ((name, srcs), strings) thy =
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   276
  let
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   277
    val atts = map (prep_attrib thy) srcs;
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   278
    val ts = map (prep_term thy) strings;
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   279
    val simps = map (fix_pat thy) ts;
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   280
  in
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   281
    (snd o PureThy.add_thmss [((name, simps), atts)]) thy
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   282
  end;
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   283
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   284
val add_fixpat = gen_add_fixpat Sign.read_term Attrib.attribute;
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   285
val add_fixpat_i = gen_add_fixpat Sign.cert_term (K I);
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   286
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   287
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   288
(*************************************************************************)
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   289
(******************************** Parsers ********************************)
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   290
(*************************************************************************)
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   291
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   292
local structure P = OuterParse and K = OuterKeyword in
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   293
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   294
val fixrec_eqn = SpecParse.opt_thm_name ":" -- P.prop;
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   295
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   296
val fixrec_strict = P.opt_keyword "permissive" >> not;
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   297
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   298
val fixrec_decl = fixrec_strict -- P.and_list1 (Scan.repeat1 fixrec_eqn);
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   299
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   300
(* this builds a parser for a new keyword, fixrec, whose functionality 
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   301
is defined by add_fixrec *)
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   302
val fixrecP =
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   303
  OuterSyntax.command "fixrec" "define recursive functions (HOLCF)" K.thy_decl
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   304
    (fixrec_decl >> (Toplevel.theory o uncurry add_fixrec));
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   305
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   306
(* fixpat parser *)
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   307
val fixpat_decl = SpecParse.opt_thm_name ":" -- Scan.repeat1 P.prop;
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   308
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   309
val fixpatP =
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   310
  OuterSyntax.command "fixpat" "define rewrites for fixrec functions" K.thy_decl
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   311
    (fixpat_decl >> (Toplevel.theory o add_fixpat));
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   312
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   313
val _ = OuterSyntax.add_parsers [fixrecP, fixpatP];
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   314
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   315
end; (* local structure *)
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   316
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   317
end; (* struct *)