src/HOLCF/fixrec_package.ML
author wenzelm
Mon, 20 Jun 2005 22:14:08 +0200
changeset 16494 6961e8ab33e1
parent 16488 38bc902946b2
child 16552 0774e9bcdb6c
permissions -rw-r--r--
added certify_prop, cert_term, cert_prop;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
16226
c17ac524d866 implementation of fixrec package
huffman
parents:
diff changeset
     1
(*  Title:      HOLCF/fixrec_package.ML
c17ac524d866 implementation of fixrec package
huffman
parents:
diff changeset
     2
    ID:         $Id$
c17ac524d866 implementation of fixrec package
huffman
parents:
diff changeset
     3
    Author:     Amber Telfer and Brian Huffman
c17ac524d866 implementation of fixrec package
huffman
parents:
diff changeset
     4
c17ac524d866 implementation of fixrec package
huffman
parents:
diff changeset
     5
Recursive function definition package for HOLCF.
c17ac524d866 implementation of fixrec package
huffman
parents:
diff changeset
     6
*)
c17ac524d866 implementation of fixrec package
huffman
parents:
diff changeset
     7
c17ac524d866 implementation of fixrec package
huffman
parents:
diff changeset
     8
signature FIXREC_PACKAGE =
c17ac524d866 implementation of fixrec package
huffman
parents:
diff changeset
     9
sig
16461
e6b431cb8e0c support theorem names and attributes for fixrec equations; also make them into simp rules by default, like primrec does
huffman
parents: 16402
diff changeset
    10
  val add_fixrec: ((string * Attrib.src list) * string) list list -> theory -> theory
16488
38bc902946b2 added add_fixrec_i, add_fixpat_i;
wenzelm
parents: 16463
diff changeset
    11
  val add_fixrec_i: ((string * theory attribute list) * term) list list -> theory -> theory
16402
36f41d5e3b3e allow theorem attributes on fixpat declarations
huffman
parents: 16401
diff changeset
    12
  val add_fixpat: ((string * Attrib.src list) * string) list -> theory -> theory
16488
38bc902946b2 added add_fixrec_i, add_fixpat_i;
wenzelm
parents: 16463
diff changeset
    13
  val add_fixpat_i: ((string * theory attribute list) * term) list -> theory -> theory
16226
c17ac524d866 implementation of fixrec package
huffman
parents:
diff changeset
    14
end;
c17ac524d866 implementation of fixrec package
huffman
parents:
diff changeset
    15
c17ac524d866 implementation of fixrec package
huffman
parents:
diff changeset
    16
structure FixrecPackage: FIXREC_PACKAGE =
c17ac524d866 implementation of fixrec package
huffman
parents:
diff changeset
    17
struct
c17ac524d866 implementation of fixrec package
huffman
parents:
diff changeset
    18
c17ac524d866 implementation of fixrec package
huffman
parents:
diff changeset
    19
(* ->> is taken from holcf_logic.ML *)
c17ac524d866 implementation of fixrec package
huffman
parents:
diff changeset
    20
(* TODO: fix dependencies so we can import HOLCFLogic here *)
c17ac524d866 implementation of fixrec package
huffman
parents:
diff changeset
    21
infixr 6 ->>;
c17ac524d866 implementation of fixrec package
huffman
parents:
diff changeset
    22
fun S ->> T = Type ("Cfun.->",[S,T]);
c17ac524d866 implementation of fixrec package
huffman
parents:
diff changeset
    23
c17ac524d866 implementation of fixrec package
huffman
parents:
diff changeset
    24
(* extern_name is taken from domain/library.ML *)
c17ac524d866 implementation of fixrec package
huffman
parents:
diff changeset
    25
fun extern_name con = case Symbol.explode con of 
c17ac524d866 implementation of fixrec package
huffman
parents:
diff changeset
    26
		   ("o"::"p"::" "::rest) => implode rest
c17ac524d866 implementation of fixrec package
huffman
parents:
diff changeset
    27
		   | _ => con;
c17ac524d866 implementation of fixrec package
huffman
parents:
diff changeset
    28
16401
57c35ede00b9 fixrec package now handles mutually-recursive definitions
huffman
parents: 16387
diff changeset
    29
val mk_trp = HOLogic.mk_Trueprop;
57c35ede00b9 fixrec package now handles mutually-recursive definitions
huffman
parents: 16387
diff changeset
    30
16226
c17ac524d866 implementation of fixrec package
huffman
parents:
diff changeset
    31
(* splits a cterm into the right and lefthand sides of equality *)
c17ac524d866 implementation of fixrec package
huffman
parents:
diff changeset
    32
fun dest_eqs (Const ("==", _)$lhs$rhs) = (lhs, rhs)
16461
e6b431cb8e0c support theorem names and attributes for fixrec equations; also make them into simp rules by default, like primrec does
huffman
parents: 16402
diff changeset
    33
  | dest_eqs (Const ("Trueprop", _)$(Const ("op =", _)$lhs$rhs)) = (lhs,rhs)
16226
c17ac524d866 implementation of fixrec package
huffman
parents:
diff changeset
    34
  | dest_eqs t = sys_error (Sign.string_of_term (sign_of (the_context())) t);
c17ac524d866 implementation of fixrec package
huffman
parents:
diff changeset
    35
16461
e6b431cb8e0c support theorem names and attributes for fixrec equations; also make them into simp rules by default, like primrec does
huffman
parents: 16402
diff changeset
    36
(* similar to Thm.head_of, but for continuous application *)
e6b431cb8e0c support theorem names and attributes for fixrec equations; also make them into simp rules by default, like primrec does
huffman
parents: 16402
diff changeset
    37
fun chead_of (Const("Cfun.Rep_CFun",_)$f$t) = chead_of f
e6b431cb8e0c support theorem names and attributes for fixrec equations; also make them into simp rules by default, like primrec does
huffman
parents: 16402
diff changeset
    38
  | chead_of u = u;
e6b431cb8e0c support theorem names and attributes for fixrec equations; also make them into simp rules by default, like primrec does
huffman
parents: 16402
diff changeset
    39
16387
67f6044c1891 cleaned up and reorganized
huffman
parents: 16226
diff changeset
    40
(* these are helpful functions copied from HOLCF/domain/library.ML *)
67f6044c1891 cleaned up and reorganized
huffman
parents: 16226
diff changeset
    41
fun %: s = Free(s,dummyT);
67f6044c1891 cleaned up and reorganized
huffman
parents: 16226
diff changeset
    42
fun %%: s = Const(s,dummyT);
67f6044c1891 cleaned up and reorganized
huffman
parents: 16226
diff changeset
    43
infix 0 ==;  fun S ==  T = %%:"==" $ S $ T;
67f6044c1891 cleaned up and reorganized
huffman
parents: 16226
diff changeset
    44
infix 1 ===; fun S === T = %%:"op =" $ S $ T;
16401
57c35ede00b9 fixrec package now handles mutually-recursive definitions
huffman
parents: 16387
diff changeset
    45
infix 9 `  ; fun f ` x = %%:"Cfun.Rep_CFun" $ f $ x;
16387
67f6044c1891 cleaned up and reorganized
huffman
parents: 16226
diff changeset
    46
67f6044c1891 cleaned up and reorganized
huffman
parents: 16226
diff changeset
    47
(* infers the type of a term *)
16402
36f41d5e3b3e allow theorem attributes on fixpat declarations
huffman
parents: 16401
diff changeset
    48
(* similar to Theory.inferT_axm, but allows any type, not just propT *)
16401
57c35ede00b9 fixrec package now handles mutually-recursive definitions
huffman
parents: 16387
diff changeset
    49
fun infer sg t =
57c35ede00b9 fixrec package now handles mutually-recursive definitions
huffman
parents: 16387
diff changeset
    50
  fst (Sign.infer_types (Sign.pp sg) sg (K NONE) (K NONE) [] true ([t],dummyT));
57c35ede00b9 fixrec package now handles mutually-recursive definitions
huffman
parents: 16387
diff changeset
    51
16402
36f41d5e3b3e allow theorem attributes on fixpat declarations
huffman
parents: 16401
diff changeset
    52
(* Similar to Term.lambda, but also allows abstraction over constants *)
16401
57c35ede00b9 fixrec package now handles mutually-recursive definitions
huffman
parents: 16387
diff changeset
    53
fun lambda' (v as Free (x, T)) t = Abs (x, T, abstract_over (v, t))
57c35ede00b9 fixrec package now handles mutually-recursive definitions
huffman
parents: 16387
diff changeset
    54
  | lambda' (v as Var ((x, _), T)) t = Abs (x, T, abstract_over (v, t))
57c35ede00b9 fixrec package now handles mutually-recursive definitions
huffman
parents: 16387
diff changeset
    55
  | lambda' (v as Const (x, T)) t = Abs (Sign.base_name x, T, abstract_over (v, t))
57c35ede00b9 fixrec package now handles mutually-recursive definitions
huffman
parents: 16387
diff changeset
    56
  | lambda' v t = raise TERM ("lambda'", [v, t]);
57c35ede00b9 fixrec package now handles mutually-recursive definitions
huffman
parents: 16387
diff changeset
    57
57c35ede00b9 fixrec package now handles mutually-recursive definitions
huffman
parents: 16387
diff changeset
    58
(* builds the expression (LAM v. rhs) *)
57c35ede00b9 fixrec package now handles mutually-recursive definitions
huffman
parents: 16387
diff changeset
    59
fun big_lambda v rhs = %%:"Cfun.Abs_CFun"$(lambda' v rhs);
57c35ede00b9 fixrec package now handles mutually-recursive definitions
huffman
parents: 16387
diff changeset
    60
57c35ede00b9 fixrec package now handles mutually-recursive definitions
huffman
parents: 16387
diff changeset
    61
(* builds the expression (LAM v1 v2 .. vn. rhs) *)
57c35ede00b9 fixrec package now handles mutually-recursive definitions
huffman
parents: 16387
diff changeset
    62
fun big_lambdas [] rhs = rhs
57c35ede00b9 fixrec package now handles mutually-recursive definitions
huffman
parents: 16387
diff changeset
    63
  | big_lambdas (v::vs) rhs = big_lambda v (big_lambdas vs rhs);
57c35ede00b9 fixrec package now handles mutually-recursive definitions
huffman
parents: 16387
diff changeset
    64
57c35ede00b9 fixrec package now handles mutually-recursive definitions
huffman
parents: 16387
diff changeset
    65
(* builds the expression (LAM <v1,v2,..,vn>. rhs) *)
57c35ede00b9 fixrec package now handles mutually-recursive definitions
huffman
parents: 16387
diff changeset
    66
fun lambda_ctuple [] rhs = big_lambda (%:"unit") rhs
57c35ede00b9 fixrec package now handles mutually-recursive definitions
huffman
parents: 16387
diff changeset
    67
  | lambda_ctuple (v::[]) rhs = big_lambda v rhs
57c35ede00b9 fixrec package now handles mutually-recursive definitions
huffman
parents: 16387
diff changeset
    68
  | lambda_ctuple (v::vs) rhs =
57c35ede00b9 fixrec package now handles mutually-recursive definitions
huffman
parents: 16387
diff changeset
    69
      %%:"Cprod.csplit"`(big_lambda v (lambda_ctuple vs rhs));
57c35ede00b9 fixrec package now handles mutually-recursive definitions
huffman
parents: 16387
diff changeset
    70
57c35ede00b9 fixrec package now handles mutually-recursive definitions
huffman
parents: 16387
diff changeset
    71
(* builds the expression <v1,v2,..,vn> *)
57c35ede00b9 fixrec package now handles mutually-recursive definitions
huffman
parents: 16387
diff changeset
    72
fun mk_ctuple [] = %%:"UU"
57c35ede00b9 fixrec package now handles mutually-recursive definitions
huffman
parents: 16387
diff changeset
    73
|   mk_ctuple (t::[]) = t
57c35ede00b9 fixrec package now handles mutually-recursive definitions
huffman
parents: 16387
diff changeset
    74
|   mk_ctuple (t::ts) = %%:"Cprod.cpair"`t`(mk_ctuple ts);
16387
67f6044c1891 cleaned up and reorganized
huffman
parents: 16226
diff changeset
    75
67f6044c1891 cleaned up and reorganized
huffman
parents: 16226
diff changeset
    76
(*************************************************************************)
16402
36f41d5e3b3e allow theorem attributes on fixpat declarations
huffman
parents: 16401
diff changeset
    77
(************* fixed-point definitions and unfolding theorems ************)
16387
67f6044c1891 cleaned up and reorganized
huffman
parents: 16226
diff changeset
    78
(*************************************************************************)
67f6044c1891 cleaned up and reorganized
huffman
parents: 16226
diff changeset
    79
16401
57c35ede00b9 fixrec package now handles mutually-recursive definitions
huffman
parents: 16387
diff changeset
    80
fun add_fixdefs eqs thy =
16226
c17ac524d866 implementation of fixrec package
huffman
parents:
diff changeset
    81
  let
16401
57c35ede00b9 fixrec package now handles mutually-recursive definitions
huffman
parents: 16387
diff changeset
    82
    val (lhss,rhss) = ListPair.unzip (map dest_eqs eqs);
57c35ede00b9 fixrec package now handles mutually-recursive definitions
huffman
parents: 16387
diff changeset
    83
    val fixpoint = %%:"Fix.fix"`lambda_ctuple lhss (mk_ctuple rhss);
57c35ede00b9 fixrec package now handles mutually-recursive definitions
huffman
parents: 16387
diff changeset
    84
    
57c35ede00b9 fixrec package now handles mutually-recursive definitions
huffman
parents: 16387
diff changeset
    85
    fun one_def (l as Const(n,T)) r =
57c35ede00b9 fixrec package now handles mutually-recursive definitions
huffman
parents: 16387
diff changeset
    86
          let val b = Sign.base_name n in (b, (b^"_fixdef", l == r)) end
57c35ede00b9 fixrec package now handles mutually-recursive definitions
huffman
parents: 16387
diff changeset
    87
      | one_def _ _ = sys_error "fixdefs: lhs not of correct form";
57c35ede00b9 fixrec package now handles mutually-recursive definitions
huffman
parents: 16387
diff changeset
    88
    fun defs [] _ = []
57c35ede00b9 fixrec package now handles mutually-recursive definitions
huffman
parents: 16387
diff changeset
    89
      | defs (l::[]) r = [one_def l r]
57c35ede00b9 fixrec package now handles mutually-recursive definitions
huffman
parents: 16387
diff changeset
    90
      | defs (l::ls) r = one_def l (%%:"Cprod.cfst"`r) :: defs ls (%%:"Cprod.csnd"`r);
57c35ede00b9 fixrec package now handles mutually-recursive definitions
huffman
parents: 16387
diff changeset
    91
    val (names, pre_fixdefs) = ListPair.unzip (defs lhss fixpoint);
57c35ede00b9 fixrec package now handles mutually-recursive definitions
huffman
parents: 16387
diff changeset
    92
    
57c35ede00b9 fixrec package now handles mutually-recursive definitions
huffman
parents: 16387
diff changeset
    93
    val fixdefs = map (inferT_axm (sign_of thy)) pre_fixdefs;
57c35ede00b9 fixrec package now handles mutually-recursive definitions
huffman
parents: 16387
diff changeset
    94
    val (thy', fixdef_thms) =
57c35ede00b9 fixrec package now handles mutually-recursive definitions
huffman
parents: 16387
diff changeset
    95
      PureThy.add_defs_i false (map Thm.no_attributes fixdefs) thy;
57c35ede00b9 fixrec package now handles mutually-recursive definitions
huffman
parents: 16387
diff changeset
    96
    val ctuple_fixdef_thm = foldr1 (fn (x,y) => cpair_equalI OF [x,y]) fixdef_thms;
57c35ede00b9 fixrec package now handles mutually-recursive definitions
huffman
parents: 16387
diff changeset
    97
    
57c35ede00b9 fixrec package now handles mutually-recursive definitions
huffman
parents: 16387
diff changeset
    98
    fun mk_cterm t = let val sg' = sign_of thy' in cterm_of sg' (infer sg' t) end;
57c35ede00b9 fixrec package now handles mutually-recursive definitions
huffman
parents: 16387
diff changeset
    99
    val ctuple_unfold_ct = mk_cterm (mk_trp (mk_ctuple lhss === mk_ctuple rhss));
57c35ede00b9 fixrec package now handles mutually-recursive definitions
huffman
parents: 16387
diff changeset
   100
    val ctuple_unfold_thm = prove_goalw_cterm [] ctuple_unfold_ct
57c35ede00b9 fixrec package now handles mutually-recursive definitions
huffman
parents: 16387
diff changeset
   101
          (fn _ => [rtac (ctuple_fixdef_thm RS fix_eq2 RS trans) 1,
57c35ede00b9 fixrec package now handles mutually-recursive definitions
huffman
parents: 16387
diff changeset
   102
                    simp_tac (simpset_of thy') 1]);
57c35ede00b9 fixrec package now handles mutually-recursive definitions
huffman
parents: 16387
diff changeset
   103
    val ctuple_induct_thm =
16461
e6b431cb8e0c support theorem names and attributes for fixrec equations; also make them into simp rules by default, like primrec does
huffman
parents: 16402
diff changeset
   104
          (space_implode "_" names ^ "_induct", ctuple_fixdef_thm RS def_fix_ind);
16401
57c35ede00b9 fixrec package now handles mutually-recursive definitions
huffman
parents: 16387
diff changeset
   105
    
57c35ede00b9 fixrec package now handles mutually-recursive definitions
huffman
parents: 16387
diff changeset
   106
    fun unfolds [] thm = []
16461
e6b431cb8e0c support theorem names and attributes for fixrec equations; also make them into simp rules by default, like primrec does
huffman
parents: 16402
diff changeset
   107
      | unfolds (n::[]) thm = [(n^"_unfold", thm)]
16401
57c35ede00b9 fixrec package now handles mutually-recursive definitions
huffman
parents: 16387
diff changeset
   108
      | unfolds (n::ns) thm = let
57c35ede00b9 fixrec package now handles mutually-recursive definitions
huffman
parents: 16387
diff changeset
   109
          val thmL = thm RS cpair_eqD1;
57c35ede00b9 fixrec package now handles mutually-recursive definitions
huffman
parents: 16387
diff changeset
   110
          val thmR = thm RS cpair_eqD2;
16461
e6b431cb8e0c support theorem names and attributes for fixrec equations; also make them into simp rules by default, like primrec does
huffman
parents: 16402
diff changeset
   111
        in (n^"_unfold", thmL) :: unfolds ns thmR end;
e6b431cb8e0c support theorem names and attributes for fixrec equations; also make them into simp rules by default, like primrec does
huffman
parents: 16402
diff changeset
   112
    val unfold_thms = unfolds names ctuple_unfold_thm;
e6b431cb8e0c support theorem names and attributes for fixrec equations; also make them into simp rules by default, like primrec does
huffman
parents: 16402
diff changeset
   113
    val thms = ctuple_induct_thm :: unfold_thms;
e6b431cb8e0c support theorem names and attributes for fixrec equations; also make them into simp rules by default, like primrec does
huffman
parents: 16402
diff changeset
   114
    val (thy'', _) = PureThy.add_thms (map Thm.no_attributes thms) thy';
16226
c17ac524d866 implementation of fixrec package
huffman
parents:
diff changeset
   115
  in
16461
e6b431cb8e0c support theorem names and attributes for fixrec equations; also make them into simp rules by default, like primrec does
huffman
parents: 16402
diff changeset
   116
    (thy'', names, fixdef_thms, map snd unfold_thms)
16401
57c35ede00b9 fixrec package now handles mutually-recursive definitions
huffman
parents: 16387
diff changeset
   117
  end;
16226
c17ac524d866 implementation of fixrec package
huffman
parents:
diff changeset
   118
16387
67f6044c1891 cleaned up and reorganized
huffman
parents: 16226
diff changeset
   119
(*************************************************************************)
67f6044c1891 cleaned up and reorganized
huffman
parents: 16226
diff changeset
   120
(*********** monadic notation and pattern matching compilation ***********)
67f6044c1891 cleaned up and reorganized
huffman
parents: 16226
diff changeset
   121
(*************************************************************************)
16226
c17ac524d866 implementation of fixrec package
huffman
parents:
diff changeset
   122
16401
57c35ede00b9 fixrec package now handles mutually-recursive definitions
huffman
parents: 16387
diff changeset
   123
fun add_names (Const(a,_), bs) = Sign.base_name a ins_string bs
57c35ede00b9 fixrec package now handles mutually-recursive definitions
huffman
parents: 16387
diff changeset
   124
  | add_names (Free(a,_) , bs) = a ins_string bs
57c35ede00b9 fixrec package now handles mutually-recursive definitions
huffman
parents: 16387
diff changeset
   125
  | add_names (f $ u     , bs) = add_names (f, add_names(u, bs))
57c35ede00b9 fixrec package now handles mutually-recursive definitions
huffman
parents: 16387
diff changeset
   126
  | add_names (Abs(a,_,t), bs) = add_names (t, a ins_string bs)
57c35ede00b9 fixrec package now handles mutually-recursive definitions
huffman
parents: 16387
diff changeset
   127
  | add_names (_         , bs) = bs;
16226
c17ac524d866 implementation of fixrec package
huffman
parents:
diff changeset
   128
c17ac524d866 implementation of fixrec package
huffman
parents:
diff changeset
   129
fun add_terms ts xs = foldr add_names xs ts;
c17ac524d866 implementation of fixrec package
huffman
parents:
diff changeset
   130
c17ac524d866 implementation of fixrec package
huffman
parents:
diff changeset
   131
(* builds a monadic term for matching a constructor pattern *)
16387
67f6044c1891 cleaned up and reorganized
huffman
parents: 16226
diff changeset
   132
fun pre_build pat rhs vs taken =
67f6044c1891 cleaned up and reorganized
huffman
parents: 16226
diff changeset
   133
  case pat of
67f6044c1891 cleaned up and reorganized
huffman
parents: 16226
diff changeset
   134
    Const("Cfun.Rep_CFun",_)$f$(v as Free(n,T)) =>
16226
c17ac524d866 implementation of fixrec package
huffman
parents:
diff changeset
   135
      pre_build f rhs (v::vs) taken
16387
67f6044c1891 cleaned up and reorganized
huffman
parents: 16226
diff changeset
   136
  | Const("Cfun.Rep_CFun",_)$f$x =>
16226
c17ac524d866 implementation of fixrec package
huffman
parents:
diff changeset
   137
      let val (rhs', v, taken') = pre_build x rhs [] taken;
16387
67f6044c1891 cleaned up and reorganized
huffman
parents: 16226
diff changeset
   138
      in pre_build f rhs' (v::vs) taken' end
67f6044c1891 cleaned up and reorganized
huffman
parents: 16226
diff changeset
   139
  | Const(c,T) =>
16226
c17ac524d866 implementation of fixrec package
huffman
parents:
diff changeset
   140
      let
c17ac524d866 implementation of fixrec package
huffman
parents:
diff changeset
   141
        val n = variant taken "v";
16387
67f6044c1891 cleaned up and reorganized
huffman
parents: 16226
diff changeset
   142
        fun result_type (Type("Cfun.->",[_,T])) (x::xs) = result_type T xs
16226
c17ac524d866 implementation of fixrec package
huffman
parents:
diff changeset
   143
          | result_type T _ = T;
c17ac524d866 implementation of fixrec package
huffman
parents:
diff changeset
   144
        val v = Free(n, result_type T vs);
16401
57c35ede00b9 fixrec package now handles mutually-recursive definitions
huffman
parents: 16387
diff changeset
   145
        val m = "match_"^(extern_name(Sign.base_name c));
57c35ede00b9 fixrec package now handles mutually-recursive definitions
huffman
parents: 16387
diff changeset
   146
        val k = lambda_ctuple vs rhs;
16226
c17ac524d866 implementation of fixrec package
huffman
parents:
diff changeset
   147
      in
16387
67f6044c1891 cleaned up and reorganized
huffman
parents: 16226
diff changeset
   148
        (%%:"Fixrec.bind"`(%%:m`v)`k, v, n::taken)
16226
c17ac524d866 implementation of fixrec package
huffman
parents:
diff changeset
   149
      end;
c17ac524d866 implementation of fixrec package
huffman
parents:
diff changeset
   150
c17ac524d866 implementation of fixrec package
huffman
parents:
diff changeset
   151
(* builds a monadic term for matching a function definition pattern *)
c17ac524d866 implementation of fixrec package
huffman
parents:
diff changeset
   152
(* returns (name, arity, matcher) *)
16387
67f6044c1891 cleaned up and reorganized
huffman
parents: 16226
diff changeset
   153
fun building pat rhs vs taken =
67f6044c1891 cleaned up and reorganized
huffman
parents: 16226
diff changeset
   154
  case pat of
67f6044c1891 cleaned up and reorganized
huffman
parents: 16226
diff changeset
   155
    Const("Cfun.Rep_CFun", _)$f$(v as Free(n,T)) =>
16226
c17ac524d866 implementation of fixrec package
huffman
parents:
diff changeset
   156
      building f rhs (v::vs) taken
16387
67f6044c1891 cleaned up and reorganized
huffman
parents: 16226
diff changeset
   157
  | Const("Cfun.Rep_CFun", _)$f$x =>
67f6044c1891 cleaned up and reorganized
huffman
parents: 16226
diff changeset
   158
      let val (rhs', v, taken') = pre_build x rhs [] taken;
67f6044c1891 cleaned up and reorganized
huffman
parents: 16226
diff changeset
   159
      in building f rhs' (v::vs) taken' end
67f6044c1891 cleaned up and reorganized
huffman
parents: 16226
diff changeset
   160
  | Const(_,_) => (pat, length vs, big_lambdas vs rhs)
67f6044c1891 cleaned up and reorganized
huffman
parents: 16226
diff changeset
   161
  | _ => sys_error "function is not declared as constant in theory";
16226
c17ac524d866 implementation of fixrec package
huffman
parents:
diff changeset
   162
16387
67f6044c1891 cleaned up and reorganized
huffman
parents: 16226
diff changeset
   163
fun match_eq eq = 
16226
c17ac524d866 implementation of fixrec package
huffman
parents:
diff changeset
   164
  let
16387
67f6044c1891 cleaned up and reorganized
huffman
parents: 16226
diff changeset
   165
    val (lhs,rhs) = dest_eqs eq;
67f6044c1891 cleaned up and reorganized
huffman
parents: 16226
diff changeset
   166
    val (Const(name,_), arity, term) =
67f6044c1891 cleaned up and reorganized
huffman
parents: 16226
diff changeset
   167
      building lhs (%%:"Fixrec.return"`rhs) [] (add_terms [eq] []);
67f6044c1891 cleaned up and reorganized
huffman
parents: 16226
diff changeset
   168
  in (name, arity, term) end;
16226
c17ac524d866 implementation of fixrec package
huffman
parents:
diff changeset
   169
c17ac524d866 implementation of fixrec package
huffman
parents:
diff changeset
   170
(* returns the sum (using +++) of the terms in ms *)
c17ac524d866 implementation of fixrec package
huffman
parents:
diff changeset
   171
(* also applies "run" to the result! *)
c17ac524d866 implementation of fixrec package
huffman
parents:
diff changeset
   172
fun fatbar arity ms =
c17ac524d866 implementation of fixrec package
huffman
parents:
diff changeset
   173
  let
c17ac524d866 implementation of fixrec package
huffman
parents:
diff changeset
   174
    fun unLAM 0 t = t
c17ac524d866 implementation of fixrec package
huffman
parents:
diff changeset
   175
      | unLAM n (_$Abs(_,_,t)) = unLAM (n-1) t
c17ac524d866 implementation of fixrec package
huffman
parents:
diff changeset
   176
      | unLAM _ _ = sys_error "FIXREC: internal error, not enough LAMs";
c17ac524d866 implementation of fixrec package
huffman
parents:
diff changeset
   177
    fun reLAM 0 t = t
c17ac524d866 implementation of fixrec package
huffman
parents:
diff changeset
   178
      | reLAM n t = reLAM (n-1) (%%:"Abs_CFun" $ Abs("",dummyT,t));
c17ac524d866 implementation of fixrec package
huffman
parents:
diff changeset
   179
    fun mplus (x,y) = %%:"Fixrec.mplus"`x`y;
c17ac524d866 implementation of fixrec package
huffman
parents:
diff changeset
   180
    val msum = foldr1 mplus (map (unLAM arity) ms);
c17ac524d866 implementation of fixrec package
huffman
parents:
diff changeset
   181
  in
c17ac524d866 implementation of fixrec package
huffman
parents:
diff changeset
   182
    reLAM arity (%%:"Fixrec.run"`msum)
c17ac524d866 implementation of fixrec package
huffman
parents:
diff changeset
   183
  end;
c17ac524d866 implementation of fixrec package
huffman
parents:
diff changeset
   184
16387
67f6044c1891 cleaned up and reorganized
huffman
parents: 16226
diff changeset
   185
fun unzip3 [] = ([],[],[])
67f6044c1891 cleaned up and reorganized
huffman
parents: 16226
diff changeset
   186
  | unzip3 ((x,y,z)::ts) =
67f6044c1891 cleaned up and reorganized
huffman
parents: 16226
diff changeset
   187
      let val (xs,ys,zs) = unzip3 ts
67f6044c1891 cleaned up and reorganized
huffman
parents: 16226
diff changeset
   188
      in (x::xs, y::ys, z::zs) end;
16226
c17ac524d866 implementation of fixrec package
huffman
parents:
diff changeset
   189
16387
67f6044c1891 cleaned up and reorganized
huffman
parents: 16226
diff changeset
   190
(* this is the pattern-matching compiler function *)
67f6044c1891 cleaned up and reorganized
huffman
parents: 16226
diff changeset
   191
fun compile_pats eqs = 
67f6044c1891 cleaned up and reorganized
huffman
parents: 16226
diff changeset
   192
  let
67f6044c1891 cleaned up and reorganized
huffman
parents: 16226
diff changeset
   193
    val ((n::names),(a::arities),mats) = unzip3 (map match_eq eqs);
67f6044c1891 cleaned up and reorganized
huffman
parents: 16226
diff changeset
   194
    val cname = if forall (fn x => n=x) names then n
67f6044c1891 cleaned up and reorganized
huffman
parents: 16226
diff changeset
   195
          else sys_error "FIXREC: all equations must define the same function";
67f6044c1891 cleaned up and reorganized
huffman
parents: 16226
diff changeset
   196
    val arity = if forall (fn x => a=x) arities then a
67f6044c1891 cleaned up and reorganized
huffman
parents: 16226
diff changeset
   197
          else sys_error "FIXREC: all equations must have the same arity";
67f6044c1891 cleaned up and reorganized
huffman
parents: 16226
diff changeset
   198
    val rhs = fatbar arity mats;
67f6044c1891 cleaned up and reorganized
huffman
parents: 16226
diff changeset
   199
  in
16401
57c35ede00b9 fixrec package now handles mutually-recursive definitions
huffman
parents: 16387
diff changeset
   200
    mk_trp (%%:cname === rhs)
16387
67f6044c1891 cleaned up and reorganized
huffman
parents: 16226
diff changeset
   201
  end;
67f6044c1891 cleaned up and reorganized
huffman
parents: 16226
diff changeset
   202
67f6044c1891 cleaned up and reorganized
huffman
parents: 16226
diff changeset
   203
(*************************************************************************)
67f6044c1891 cleaned up and reorganized
huffman
parents: 16226
diff changeset
   204
(********************** Proving associated theorems **********************)
67f6044c1891 cleaned up and reorganized
huffman
parents: 16226
diff changeset
   205
(*************************************************************************)
67f6044c1891 cleaned up and reorganized
huffman
parents: 16226
diff changeset
   206
16461
e6b431cb8e0c support theorem names and attributes for fixrec equations; also make them into simp rules by default, like primrec does
huffman
parents: 16402
diff changeset
   207
fun prove_simp thy unfold_thm t =
16226
c17ac524d866 implementation of fixrec package
huffman
parents:
diff changeset
   208
  let
c17ac524d866 implementation of fixrec package
huffman
parents:
diff changeset
   209
    val ss = simpset_of thy;
16461
e6b431cb8e0c support theorem names and attributes for fixrec equations; also make them into simp rules by default, like primrec does
huffman
parents: 16402
diff changeset
   210
    val ct = cterm_of thy t;
16387
67f6044c1891 cleaned up and reorganized
huffman
parents: 16226
diff changeset
   211
    val thm = prove_goalw_cterm [] ct
16463
342d74ca8815 fixrec shows unsolved subgoals when proofs of rewrites fail
huffman
parents: 16461
diff changeset
   212
      (fn _ => [rtac (unfold_thm RS ssubst_lhs) 1, simp_tac ss 1]);
16226
c17ac524d866 implementation of fixrec package
huffman
parents:
diff changeset
   213
  in thm end;
c17ac524d866 implementation of fixrec package
huffman
parents:
diff changeset
   214
c17ac524d866 implementation of fixrec package
huffman
parents:
diff changeset
   215
(* this proves that each equation is a theorem *)
16461
e6b431cb8e0c support theorem names and attributes for fixrec equations; also make them into simp rules by default, like primrec does
huffman
parents: 16402
diff changeset
   216
fun prove_simps thy (unfold_thm,ts) = map (prove_simp thy unfold_thm) ts;
16401
57c35ede00b9 fixrec package now handles mutually-recursive definitions
huffman
parents: 16387
diff changeset
   217
57c35ede00b9 fixrec package now handles mutually-recursive definitions
huffman
parents: 16387
diff changeset
   218
(* proves the pattern matching equations as theorems, using unfold *)
16461
e6b431cb8e0c support theorem names and attributes for fixrec equations; also make them into simp rules by default, like primrec does
huffman
parents: 16402
diff changeset
   219
fun make_simps cnames unfold_thms namess attss tss thy = 
16401
57c35ede00b9 fixrec package now handles mutually-recursive definitions
huffman
parents: 16387
diff changeset
   220
  let
16461
e6b431cb8e0c support theorem names and attributes for fixrec equations; also make them into simp rules by default, like primrec does
huffman
parents: 16402
diff changeset
   221
    val thm_names = map (fn name => name^"_simps") cnames;
e6b431cb8e0c support theorem names and attributes for fixrec equations; also make them into simp rules by default, like primrec does
huffman
parents: 16402
diff changeset
   222
    val rew_thmss = map (prove_simps thy) (unfold_thms ~~ tss);
e6b431cb8e0c support theorem names and attributes for fixrec equations; also make them into simp rules by default, like primrec does
huffman
parents: 16402
diff changeset
   223
    val thms = (List.concat namess ~~ List.concat rew_thmss) ~~ List.concat attss;
e6b431cb8e0c support theorem names and attributes for fixrec equations; also make them into simp rules by default, like primrec does
huffman
parents: 16402
diff changeset
   224
    val (thy',_) = PureThy.add_thms thms thy;
e6b431cb8e0c support theorem names and attributes for fixrec equations; also make them into simp rules by default, like primrec does
huffman
parents: 16402
diff changeset
   225
    val thmss = thm_names ~~ rew_thmss;
e6b431cb8e0c support theorem names and attributes for fixrec equations; also make them into simp rules by default, like primrec does
huffman
parents: 16402
diff changeset
   226
    val simp_attribute = rpair [Simplifier.simp_add_global];
16401
57c35ede00b9 fixrec package now handles mutually-recursive definitions
huffman
parents: 16387
diff changeset
   227
  in
16461
e6b431cb8e0c support theorem names and attributes for fixrec equations; also make them into simp rules by default, like primrec does
huffman
parents: 16402
diff changeset
   228
    (#1 o PureThy.add_thmss (map simp_attribute thmss)) thy'
16401
57c35ede00b9 fixrec package now handles mutually-recursive definitions
huffman
parents: 16387
diff changeset
   229
  end;
16226
c17ac524d866 implementation of fixrec package
huffman
parents:
diff changeset
   230
16387
67f6044c1891 cleaned up and reorganized
huffman
parents: 16226
diff changeset
   231
(*************************************************************************)
67f6044c1891 cleaned up and reorganized
huffman
parents: 16226
diff changeset
   232
(************************* Main fixrec function **************************)
67f6044c1891 cleaned up and reorganized
huffman
parents: 16226
diff changeset
   233
(*************************************************************************)
16226
c17ac524d866 implementation of fixrec package
huffman
parents:
diff changeset
   234
16387
67f6044c1891 cleaned up and reorganized
huffman
parents: 16226
diff changeset
   235
(* this calls the main processing function and then returns the new state *)
16488
38bc902946b2 added add_fixrec_i, add_fixpat_i;
wenzelm
parents: 16463
diff changeset
   236
fun gen_add_fixrec prep_prop prep_attrib blocks thy =
16226
c17ac524d866 implementation of fixrec package
huffman
parents:
diff changeset
   237
  let
16461
e6b431cb8e0c support theorem names and attributes for fixrec equations; also make them into simp rules by default, like primrec does
huffman
parents: 16402
diff changeset
   238
    fun split_list2 xss = split_list (map split_list xss);
e6b431cb8e0c support theorem names and attributes for fixrec equations; also make them into simp rules by default, like primrec does
huffman
parents: 16402
diff changeset
   239
    val ((namess, srcsss), strss) = apfst split_list2 (split_list2 blocks);
16488
38bc902946b2 added add_fixrec_i, add_fixpat_i;
wenzelm
parents: 16463
diff changeset
   240
    val attss = map (map (map (prep_attrib thy))) srcsss;
38bc902946b2 added add_fixrec_i, add_fixpat_i;
wenzelm
parents: 16463
diff changeset
   241
    val tss = map (map (prep_prop thy)) strss;
38bc902946b2 added add_fixrec_i, add_fixpat_i;
wenzelm
parents: 16463
diff changeset
   242
    val ts' = map (infer thy o compile_pats) tss;
16461
e6b431cb8e0c support theorem names and attributes for fixrec equations; also make them into simp rules by default, like primrec does
huffman
parents: 16402
diff changeset
   243
    val (thy', cnames, fixdef_thms, unfold_thms) = add_fixdefs ts' thy;
16226
c17ac524d866 implementation of fixrec package
huffman
parents:
diff changeset
   244
  in
16461
e6b431cb8e0c support theorem names and attributes for fixrec equations; also make them into simp rules by default, like primrec does
huffman
parents: 16402
diff changeset
   245
    make_simps cnames unfold_thms namess attss tss thy'
16226
c17ac524d866 implementation of fixrec package
huffman
parents:
diff changeset
   246
  end;
c17ac524d866 implementation of fixrec package
huffman
parents:
diff changeset
   247
16488
38bc902946b2 added add_fixrec_i, add_fixpat_i;
wenzelm
parents: 16463
diff changeset
   248
val add_fixrec = gen_add_fixrec Sign.read_prop Attrib.global_attribute;
38bc902946b2 added add_fixrec_i, add_fixpat_i;
wenzelm
parents: 16463
diff changeset
   249
val add_fixrec_i = gen_add_fixrec Sign.cert_prop (K I);
38bc902946b2 added add_fixrec_i, add_fixpat_i;
wenzelm
parents: 16463
diff changeset
   250
38bc902946b2 added add_fixrec_i, add_fixpat_i;
wenzelm
parents: 16463
diff changeset
   251
16387
67f6044c1891 cleaned up and reorganized
huffman
parents: 16226
diff changeset
   252
(*************************************************************************)
67f6044c1891 cleaned up and reorganized
huffman
parents: 16226
diff changeset
   253
(******************************** Fixpat *********************************)
67f6044c1891 cleaned up and reorganized
huffman
parents: 16226
diff changeset
   254
(*************************************************************************)
16226
c17ac524d866 implementation of fixrec package
huffman
parents:
diff changeset
   255
16488
38bc902946b2 added add_fixrec_i, add_fixpat_i;
wenzelm
parents: 16463
diff changeset
   256
fun fix_pat prep_term thy pat = 
16226
c17ac524d866 implementation of fixrec package
huffman
parents:
diff changeset
   257
  let
16488
38bc902946b2 added add_fixrec_i, add_fixpat_i;
wenzelm
parents: 16463
diff changeset
   258
    val t = prep_term thy pat;
16387
67f6044c1891 cleaned up and reorganized
huffman
parents: 16226
diff changeset
   259
    val T = fastype_of t;
16401
57c35ede00b9 fixrec package now handles mutually-recursive definitions
huffman
parents: 16387
diff changeset
   260
    val eq = mk_trp (HOLogic.eq_const T $ t $ Var (("x",0),T));
16461
e6b431cb8e0c support theorem names and attributes for fixrec equations; also make them into simp rules by default, like primrec does
huffman
parents: 16402
diff changeset
   261
    val cname = case chead_of t of Const(c,_) => c | _ =>
e6b431cb8e0c support theorem names and attributes for fixrec equations; also make them into simp rules by default, like primrec does
huffman
parents: 16402
diff changeset
   262
              sys_error "FIXPAT: function is not declared as constant in theory";
16488
38bc902946b2 added add_fixrec_i, add_fixpat_i;
wenzelm
parents: 16463
diff changeset
   263
    val unfold_thm = PureThy.get_thm thy (Name (cname^"_unfold"));
38bc902946b2 added add_fixrec_i, add_fixpat_i;
wenzelm
parents: 16463
diff changeset
   264
    val rew = prove_goalw_cterm [] (cterm_of thy eq)
16387
67f6044c1891 cleaned up and reorganized
huffman
parents: 16226
diff changeset
   265
          (fn _ => [stac unfold_thm 1, simp_tac (simpset_of thy) 1]);
16402
36f41d5e3b3e allow theorem attributes on fixpat declarations
huffman
parents: 16401
diff changeset
   266
  in rew end;
16226
c17ac524d866 implementation of fixrec package
huffman
parents:
diff changeset
   267
16488
38bc902946b2 added add_fixrec_i, add_fixpat_i;
wenzelm
parents: 16463
diff changeset
   268
fun gen_add_fixpat prep_term prep_attrib pats thy =
16402
36f41d5e3b3e allow theorem attributes on fixpat declarations
huffman
parents: 16401
diff changeset
   269
  let
36f41d5e3b3e allow theorem attributes on fixpat declarations
huffman
parents: 16401
diff changeset
   270
    val ((names, srcss), strings) = apfst ListPair.unzip (ListPair.unzip pats);
16488
38bc902946b2 added add_fixrec_i, add_fixpat_i;
wenzelm
parents: 16463
diff changeset
   271
    val atts = map (map (prep_attrib thy)) srcss;
38bc902946b2 added add_fixrec_i, add_fixpat_i;
wenzelm
parents: 16463
diff changeset
   272
    val simps = map (fix_pat prep_term thy) strings;
16461
e6b431cb8e0c support theorem names and attributes for fixrec equations; also make them into simp rules by default, like primrec does
huffman
parents: 16402
diff changeset
   273
    val (thy', _) = PureThy.add_thms ((names ~~ simps) ~~ atts) thy;
16402
36f41d5e3b3e allow theorem attributes on fixpat declarations
huffman
parents: 16401
diff changeset
   274
  in thy' end;
16226
c17ac524d866 implementation of fixrec package
huffman
parents:
diff changeset
   275
16488
38bc902946b2 added add_fixrec_i, add_fixpat_i;
wenzelm
parents: 16463
diff changeset
   276
val add_fixpat = gen_add_fixpat Sign.read_term Attrib.global_attribute;
38bc902946b2 added add_fixrec_i, add_fixpat_i;
wenzelm
parents: 16463
diff changeset
   277
val add_fixpat_i = gen_add_fixpat Sign.cert_term (K I);
38bc902946b2 added add_fixrec_i, add_fixpat_i;
wenzelm
parents: 16463
diff changeset
   278
38bc902946b2 added add_fixrec_i, add_fixpat_i;
wenzelm
parents: 16463
diff changeset
   279
16387
67f6044c1891 cleaned up and reorganized
huffman
parents: 16226
diff changeset
   280
(*************************************************************************)
67f6044c1891 cleaned up and reorganized
huffman
parents: 16226
diff changeset
   281
(******************************** Parsers ********************************)
67f6044c1891 cleaned up and reorganized
huffman
parents: 16226
diff changeset
   282
(*************************************************************************)
16226
c17ac524d866 implementation of fixrec package
huffman
parents:
diff changeset
   283
c17ac524d866 implementation of fixrec package
huffman
parents:
diff changeset
   284
local structure P = OuterParse and K = OuterSyntax.Keyword in
c17ac524d866 implementation of fixrec package
huffman
parents:
diff changeset
   285
16461
e6b431cb8e0c support theorem names and attributes for fixrec equations; also make them into simp rules by default, like primrec does
huffman
parents: 16402
diff changeset
   286
val fixrec_decl = P.and_list1 (Scan.repeat1 (P.opt_thm_name ":" -- P.prop));
16226
c17ac524d866 implementation of fixrec package
huffman
parents:
diff changeset
   287
c17ac524d866 implementation of fixrec package
huffman
parents:
diff changeset
   288
(* this builds a parser for a new keyword, fixrec, whose functionality 
c17ac524d866 implementation of fixrec package
huffman
parents:
diff changeset
   289
is defined by add_fixrec *)
c17ac524d866 implementation of fixrec package
huffman
parents:
diff changeset
   290
val fixrecP =
16402
36f41d5e3b3e allow theorem attributes on fixpat declarations
huffman
parents: 16401
diff changeset
   291
  OuterSyntax.command "fixrec" "define recursive functions (HOLCF)" K.thy_decl
16226
c17ac524d866 implementation of fixrec package
huffman
parents:
diff changeset
   292
    (fixrec_decl >> (Toplevel.theory o add_fixrec));
c17ac524d866 implementation of fixrec package
huffman
parents:
diff changeset
   293
c17ac524d866 implementation of fixrec package
huffman
parents:
diff changeset
   294
(* this adds the parser for fixrec to the syntax *)
c17ac524d866 implementation of fixrec package
huffman
parents:
diff changeset
   295
val _ = OuterSyntax.add_parsers [fixrecP];
c17ac524d866 implementation of fixrec package
huffman
parents:
diff changeset
   296
c17ac524d866 implementation of fixrec package
huffman
parents:
diff changeset
   297
(* fixpat parser *)
16461
e6b431cb8e0c support theorem names and attributes for fixrec equations; also make them into simp rules by default, like primrec does
huffman
parents: 16402
diff changeset
   298
val fixpat_decl = Scan.repeat1 (P.opt_thm_name ":" -- P.prop);
16226
c17ac524d866 implementation of fixrec package
huffman
parents:
diff changeset
   299
c17ac524d866 implementation of fixrec package
huffman
parents:
diff changeset
   300
val fixpatP =
16402
36f41d5e3b3e allow theorem attributes on fixpat declarations
huffman
parents: 16401
diff changeset
   301
  OuterSyntax.command "fixpat" "define rewrites for fixrec functions" K.thy_decl
16226
c17ac524d866 implementation of fixrec package
huffman
parents:
diff changeset
   302
    (fixpat_decl >> (Toplevel.theory o add_fixpat));
c17ac524d866 implementation of fixrec package
huffman
parents:
diff changeset
   303
c17ac524d866 implementation of fixrec package
huffman
parents:
diff changeset
   304
val _ = OuterSyntax.add_parsers [fixpatP];
c17ac524d866 implementation of fixrec package
huffman
parents:
diff changeset
   305
c17ac524d866 implementation of fixrec package
huffman
parents:
diff changeset
   306
end; (* local structure *)
c17ac524d866 implementation of fixrec package
huffman
parents:
diff changeset
   307
c17ac524d866 implementation of fixrec package
huffman
parents:
diff changeset
   308
end; (* struct *)