src/HOLCF/Tools/Domain/domain_isomorphism.ML
author huffman
Thu, 19 Nov 2009 12:05:49 -0800
changeset 33792 002e0e017311
parent 33791 fef59343b4b3
child 33793 5beafabffa07
permissions -rw-r--r--
change Theory.requires
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
33774
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
     1
(*  Title:      HOLCF/Tools/domain/domain_isomorphism.ML
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
     2
    Author:     Brian Huffman
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
     3
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
     4
Defines new types satisfying the given domain equations.
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
     5
*)
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
     6
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
     7
signature DOMAIN_ISOMORPHISM =
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
     8
sig
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
     9
  val domain_isomorphism:
33791
fef59343b4b3 use theory data for REP_simps and isodefl_rules
huffman
parents: 33790
diff changeset
    10
    (string list * binding * mixfix * typ) list -> theory -> theory
33774
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
    11
  val domain_isomorphism_cmd:
33791
fef59343b4b3 use theory data for REP_simps and isodefl_rules
huffman
parents: 33790
diff changeset
    12
    (string list * binding * mixfix * string) list -> theory -> theory
fef59343b4b3 use theory data for REP_simps and isodefl_rules
huffman
parents: 33790
diff changeset
    13
  val add_type_constructor:
fef59343b4b3 use theory data for REP_simps and isodefl_rules
huffman
parents: 33790
diff changeset
    14
    (string * term * string * thm  * thm) -> theory -> theory
33774
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
    15
end;
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
    16
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
    17
structure Domain_Isomorphism :> DOMAIN_ISOMORPHISM =
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
    18
struct
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
    19
33775
7a1518c42c56 cleaned up; factored out fixed-point definition code
huffman
parents: 33774
diff changeset
    20
val beta_ss =
7a1518c42c56 cleaned up; factored out fixed-point definition code
huffman
parents: 33774
diff changeset
    21
  HOL_basic_ss
7a1518c42c56 cleaned up; factored out fixed-point definition code
huffman
parents: 33774
diff changeset
    22
    addsimps simp_thms
7a1518c42c56 cleaned up; factored out fixed-point definition code
huffman
parents: 33774
diff changeset
    23
    addsimps [@{thm beta_cfun}]
7a1518c42c56 cleaned up; factored out fixed-point definition code
huffman
parents: 33774
diff changeset
    24
    addsimprocs [@{simproc cont_proc}];
7a1518c42c56 cleaned up; factored out fixed-point definition code
huffman
parents: 33774
diff changeset
    25
7a1518c42c56 cleaned up; factored out fixed-point definition code
huffman
parents: 33774
diff changeset
    26
val beta_tac = simp_tac beta_ss;
7a1518c42c56 cleaned up; factored out fixed-point definition code
huffman
parents: 33774
diff changeset
    27
33774
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
    28
(******************************************************************************)
33791
fef59343b4b3 use theory data for REP_simps and isodefl_rules
huffman
parents: 33790
diff changeset
    29
(******************************** theory data *********************************)
fef59343b4b3 use theory data for REP_simps and isodefl_rules
huffman
parents: 33790
diff changeset
    30
(******************************************************************************)
fef59343b4b3 use theory data for REP_simps and isodefl_rules
huffman
parents: 33790
diff changeset
    31
fef59343b4b3 use theory data for REP_simps and isodefl_rules
huffman
parents: 33790
diff changeset
    32
structure DeflData = Theory_Data
fef59343b4b3 use theory data for REP_simps and isodefl_rules
huffman
parents: 33790
diff changeset
    33
(
fef59343b4b3 use theory data for REP_simps and isodefl_rules
huffman
parents: 33790
diff changeset
    34
  type T = term Symtab.table;
fef59343b4b3 use theory data for REP_simps and isodefl_rules
huffman
parents: 33790
diff changeset
    35
  val empty = Symtab.empty;
fef59343b4b3 use theory data for REP_simps and isodefl_rules
huffman
parents: 33790
diff changeset
    36
  val extend = I;
fef59343b4b3 use theory data for REP_simps and isodefl_rules
huffman
parents: 33790
diff changeset
    37
  fun merge data = Symtab.merge (K true) data;
fef59343b4b3 use theory data for REP_simps and isodefl_rules
huffman
parents: 33790
diff changeset
    38
);
fef59343b4b3 use theory data for REP_simps and isodefl_rules
huffman
parents: 33790
diff changeset
    39
fef59343b4b3 use theory data for REP_simps and isodefl_rules
huffman
parents: 33790
diff changeset
    40
structure MapData = Theory_Data
fef59343b4b3 use theory data for REP_simps and isodefl_rules
huffman
parents: 33790
diff changeset
    41
(
fef59343b4b3 use theory data for REP_simps and isodefl_rules
huffman
parents: 33790
diff changeset
    42
  type T = string Symtab.table;
fef59343b4b3 use theory data for REP_simps and isodefl_rules
huffman
parents: 33790
diff changeset
    43
  val empty = Symtab.empty;
fef59343b4b3 use theory data for REP_simps and isodefl_rules
huffman
parents: 33790
diff changeset
    44
  val extend = I;
fef59343b4b3 use theory data for REP_simps and isodefl_rules
huffman
parents: 33790
diff changeset
    45
  fun merge data = Symtab.merge (K true) data;
fef59343b4b3 use theory data for REP_simps and isodefl_rules
huffman
parents: 33790
diff changeset
    46
);
fef59343b4b3 use theory data for REP_simps and isodefl_rules
huffman
parents: 33790
diff changeset
    47
fef59343b4b3 use theory data for REP_simps and isodefl_rules
huffman
parents: 33790
diff changeset
    48
structure RepData = Theory_Data
fef59343b4b3 use theory data for REP_simps and isodefl_rules
huffman
parents: 33790
diff changeset
    49
(
fef59343b4b3 use theory data for REP_simps and isodefl_rules
huffman
parents: 33790
diff changeset
    50
  type T = thm list;
fef59343b4b3 use theory data for REP_simps and isodefl_rules
huffman
parents: 33790
diff changeset
    51
  val empty = [];
fef59343b4b3 use theory data for REP_simps and isodefl_rules
huffman
parents: 33790
diff changeset
    52
  val extend = I;
fef59343b4b3 use theory data for REP_simps and isodefl_rules
huffman
parents: 33790
diff changeset
    53
  val merge = Thm.merge_thms;
fef59343b4b3 use theory data for REP_simps and isodefl_rules
huffman
parents: 33790
diff changeset
    54
);
fef59343b4b3 use theory data for REP_simps and isodefl_rules
huffman
parents: 33790
diff changeset
    55
fef59343b4b3 use theory data for REP_simps and isodefl_rules
huffman
parents: 33790
diff changeset
    56
structure IsodeflData = Theory_Data
fef59343b4b3 use theory data for REP_simps and isodefl_rules
huffman
parents: 33790
diff changeset
    57
(
fef59343b4b3 use theory data for REP_simps and isodefl_rules
huffman
parents: 33790
diff changeset
    58
  type T = thm list;
fef59343b4b3 use theory data for REP_simps and isodefl_rules
huffman
parents: 33790
diff changeset
    59
  val empty = [];
fef59343b4b3 use theory data for REP_simps and isodefl_rules
huffman
parents: 33790
diff changeset
    60
  val extend = I;
fef59343b4b3 use theory data for REP_simps and isodefl_rules
huffman
parents: 33790
diff changeset
    61
  val merge = Thm.merge_thms;
fef59343b4b3 use theory data for REP_simps and isodefl_rules
huffman
parents: 33790
diff changeset
    62
);
fef59343b4b3 use theory data for REP_simps and isodefl_rules
huffman
parents: 33790
diff changeset
    63
fef59343b4b3 use theory data for REP_simps and isodefl_rules
huffman
parents: 33790
diff changeset
    64
fun add_type_constructor
fef59343b4b3 use theory data for REP_simps and isodefl_rules
huffman
parents: 33790
diff changeset
    65
  (tname, defl_const, map_name, REP_thm, isodefl_thm) =
fef59343b4b3 use theory data for REP_simps and isodefl_rules
huffman
parents: 33790
diff changeset
    66
    DeflData.map (Symtab.insert (K true) (tname, defl_const))
fef59343b4b3 use theory data for REP_simps and isodefl_rules
huffman
parents: 33790
diff changeset
    67
    #> MapData.map (Symtab.insert (K true) (tname, map_name))
fef59343b4b3 use theory data for REP_simps and isodefl_rules
huffman
parents: 33790
diff changeset
    68
    #> RepData.map (Thm.add_thm REP_thm)
fef59343b4b3 use theory data for REP_simps and isodefl_rules
huffman
parents: 33790
diff changeset
    69
    #> IsodeflData.map (Thm.add_thm isodefl_thm);
fef59343b4b3 use theory data for REP_simps and isodefl_rules
huffman
parents: 33790
diff changeset
    70
fef59343b4b3 use theory data for REP_simps and isodefl_rules
huffman
parents: 33790
diff changeset
    71
fef59343b4b3 use theory data for REP_simps and isodefl_rules
huffman
parents: 33790
diff changeset
    72
(******************************************************************************)
33774
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
    73
(******************************* building types *******************************)
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
    74
(******************************************************************************)
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
    75
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
    76
(* ->> is taken from holcf_logic.ML *)
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
    77
fun cfunT (T, U) = Type(@{type_name "->"}, [T, U]);
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
    78
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
    79
infixr 6 ->>; val (op ->>) = cfunT;
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
    80
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
    81
fun dest_cfunT (Type(@{type_name "->"}, [T, U])) = (T, U)
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
    82
  | dest_cfunT T = raise TYPE ("dest_cfunT", [T], []);
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
    83
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
    84
fun tupleT [] = HOLogic.unitT
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
    85
  | tupleT [T] = T
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
    86
  | tupleT (T :: Ts) = HOLogic.mk_prodT (T, tupleT Ts);
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
    87
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
    88
val deflT = @{typ "udom alg_defl"};
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
    89
33785
2f2d9eb37084 automate definition of map functions; remove unused code
huffman
parents: 33784
diff changeset
    90
fun mapT (T as Type (_, Ts)) =
2f2d9eb37084 automate definition of map functions; remove unused code
huffman
parents: 33784
diff changeset
    91
  Library.foldr cfunT (map (fn T => T ->> T) Ts, T ->> T);     
2f2d9eb37084 automate definition of map functions; remove unused code
huffman
parents: 33784
diff changeset
    92
33774
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
    93
(******************************************************************************)
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
    94
(******************************* building terms *******************************)
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
    95
(******************************************************************************)
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
    96
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
    97
(* builds the expression (v1,v2,..,vn) *)
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
    98
fun mk_tuple [] = HOLogic.unit
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
    99
|   mk_tuple (t::[]) = t
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   100
|   mk_tuple (t::ts) = HOLogic.mk_prod (t, mk_tuple ts);
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   101
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   102
(* builds the expression (%(v1,v2,..,vn). rhs) *)
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   103
fun lambda_tuple [] rhs = Term.lambda (Free("unit", HOLogic.unitT)) rhs
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   104
  | lambda_tuple (v::[]) rhs = Term.lambda v rhs
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   105
  | lambda_tuple (v::vs) rhs =
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   106
      HOLogic.mk_split (Term.lambda v (lambda_tuple vs rhs));
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   107
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   108
(* continuous application and abstraction *)
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   109
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   110
fun capply_const (S, T) =
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   111
  Const(@{const_name Rep_CFun}, (S ->> T) --> (S --> T));
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   112
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   113
fun cabs_const (S, T) =
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   114
  Const(@{const_name Abs_CFun}, (S --> T) --> (S ->> T));
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   115
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   116
fun mk_cabs t =
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   117
  let val T = Term.fastype_of t
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   118
  in cabs_const (Term.domain_type T, Term.range_type T) $ t end
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   119
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   120
(* builds the expression (LAM v. rhs) *)
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   121
fun big_lambda v rhs =
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   122
  cabs_const (Term.fastype_of v, Term.fastype_of rhs) $ Term.lambda v rhs;
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   123
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   124
(* builds the expression (LAM v1 v2 .. vn. rhs) *)
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   125
fun big_lambdas [] rhs = rhs
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   126
  | big_lambdas (v::vs) rhs = big_lambda v (big_lambdas vs rhs);
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   127
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   128
fun mk_capply (t, u) =
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   129
  let val (S, T) =
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   130
    case Term.fastype_of t of
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   131
        Type(@{type_name "->"}, [S, T]) => (S, T)
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   132
      | _ => raise TERM ("mk_capply " ^ ML_Syntax.print_list ML_Syntax.print_term [t, u], [t, u]);
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   133
  in capply_const (S, T) $ t $ u end;
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   134
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   135
(* miscellaneous term constructions *)
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   136
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   137
val mk_trp = HOLogic.mk_Trueprop;
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   138
33775
7a1518c42c56 cleaned up; factored out fixed-point definition code
huffman
parents: 33774
diff changeset
   139
val mk_fst = HOLogic.mk_fst;
7a1518c42c56 cleaned up; factored out fixed-point definition code
huffman
parents: 33774
diff changeset
   140
val mk_snd = HOLogic.mk_snd;
7a1518c42c56 cleaned up; factored out fixed-point definition code
huffman
parents: 33774
diff changeset
   141
33774
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   142
fun mk_cont t =
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   143
  let val T = Term.fastype_of t
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   144
  in Const(@{const_name cont}, T --> HOLogic.boolT) $ t end;
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   145
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   146
fun mk_fix t =
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   147
  let val (T, _) = dest_cfunT (Term.fastype_of t)
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   148
  in mk_capply (Const(@{const_name fix}, (T ->> T) ->> T), t) end;
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   149
33785
2f2d9eb37084 automate definition of map functions; remove unused code
huffman
parents: 33784
diff changeset
   150
fun ID_const T = Const (@{const_name ID}, cfunT (T, T));
2f2d9eb37084 automate definition of map functions; remove unused code
huffman
parents: 33784
diff changeset
   151
2f2d9eb37084 automate definition of map functions; remove unused code
huffman
parents: 33784
diff changeset
   152
fun cfcomp_const (T, U, V) =
2f2d9eb37084 automate definition of map functions; remove unused code
huffman
parents: 33784
diff changeset
   153
  Const (@{const_name cfcomp}, (U ->> V) ->> (T ->> U) ->> (T ->> V));
2f2d9eb37084 automate definition of map functions; remove unused code
huffman
parents: 33784
diff changeset
   154
2f2d9eb37084 automate definition of map functions; remove unused code
huffman
parents: 33784
diff changeset
   155
fun mk_cfcomp (f, g) =
2f2d9eb37084 automate definition of map functions; remove unused code
huffman
parents: 33784
diff changeset
   156
  let
2f2d9eb37084 automate definition of map functions; remove unused code
huffman
parents: 33784
diff changeset
   157
    val (U, V) = dest_cfunT (Term.fastype_of f);
2f2d9eb37084 automate definition of map functions; remove unused code
huffman
parents: 33784
diff changeset
   158
    val (T, U') = dest_cfunT (Term.fastype_of g);
2f2d9eb37084 automate definition of map functions; remove unused code
huffman
parents: 33784
diff changeset
   159
  in
2f2d9eb37084 automate definition of map functions; remove unused code
huffman
parents: 33784
diff changeset
   160
    if U = U'
2f2d9eb37084 automate definition of map functions; remove unused code
huffman
parents: 33784
diff changeset
   161
    then mk_capply (mk_capply (cfcomp_const (T, U, V), f), g)
2f2d9eb37084 automate definition of map functions; remove unused code
huffman
parents: 33784
diff changeset
   162
    else raise TYPE ("mk_cfcomp", [U, U'], [f, g])
2f2d9eb37084 automate definition of map functions; remove unused code
huffman
parents: 33784
diff changeset
   163
  end;
2f2d9eb37084 automate definition of map functions; remove unused code
huffman
parents: 33784
diff changeset
   164
33774
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   165
fun mk_Rep_of T =
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   166
  Const (@{const_name Rep_of}, Term.itselfT T --> deflT) $ Logic.mk_type T;
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   167
33778
9121ea165a40 automate definition of rep/abs functions
huffman
parents: 33777
diff changeset
   168
fun coerce_const T = Const (@{const_name coerce}, T);
9121ea165a40 automate definition of rep/abs functions
huffman
parents: 33777
diff changeset
   169
33788
481bc899febf automate isodefl proof
huffman
parents: 33785
diff changeset
   170
fun isodefl_const T =
481bc899febf automate isodefl proof
huffman
parents: 33785
diff changeset
   171
  Const (@{const_name isodefl}, (T ->> T) --> deflT --> HOLogic.boolT);
481bc899febf automate isodefl proof
huffman
parents: 33785
diff changeset
   172
33775
7a1518c42c56 cleaned up; factored out fixed-point definition code
huffman
parents: 33774
diff changeset
   173
(* splits a cterm into the right and lefthand sides of equality *)
7a1518c42c56 cleaned up; factored out fixed-point definition code
huffman
parents: 33774
diff changeset
   174
fun dest_eqs t = HOLogic.dest_eq (HOLogic.dest_Trueprop t);
7a1518c42c56 cleaned up; factored out fixed-point definition code
huffman
parents: 33774
diff changeset
   175
7a1518c42c56 cleaned up; factored out fixed-point definition code
huffman
parents: 33774
diff changeset
   176
fun mk_eqs (t, u) = HOLogic.mk_Trueprop (HOLogic.mk_eq (t, u));
7a1518c42c56 cleaned up; factored out fixed-point definition code
huffman
parents: 33774
diff changeset
   177
7a1518c42c56 cleaned up; factored out fixed-point definition code
huffman
parents: 33774
diff changeset
   178
(******************************************************************************)
7a1518c42c56 cleaned up; factored out fixed-point definition code
huffman
parents: 33774
diff changeset
   179
(*************** fixed-point definitions and unfolding theorems ***************)
7a1518c42c56 cleaned up; factored out fixed-point definition code
huffman
parents: 33774
diff changeset
   180
(******************************************************************************)
7a1518c42c56 cleaned up; factored out fixed-point definition code
huffman
parents: 33774
diff changeset
   181
7a1518c42c56 cleaned up; factored out fixed-point definition code
huffman
parents: 33774
diff changeset
   182
fun add_fixdefs
7a1518c42c56 cleaned up; factored out fixed-point definition code
huffman
parents: 33774
diff changeset
   183
    (spec : (binding * term) list)
33788
481bc899febf automate isodefl proof
huffman
parents: 33785
diff changeset
   184
    (thy : theory) : (thm list * thm list) * theory =
33775
7a1518c42c56 cleaned up; factored out fixed-point definition code
huffman
parents: 33774
diff changeset
   185
  let
7a1518c42c56 cleaned up; factored out fixed-point definition code
huffman
parents: 33774
diff changeset
   186
    val binds = map fst spec;
7a1518c42c56 cleaned up; factored out fixed-point definition code
huffman
parents: 33774
diff changeset
   187
    val (lhss, rhss) = ListPair.unzip (map (dest_eqs o snd) spec);
7a1518c42c56 cleaned up; factored out fixed-point definition code
huffman
parents: 33774
diff changeset
   188
    val functional = lambda_tuple lhss (mk_tuple rhss);
7a1518c42c56 cleaned up; factored out fixed-point definition code
huffman
parents: 33774
diff changeset
   189
    val fixpoint = mk_fix (mk_cabs functional);
7a1518c42c56 cleaned up; factored out fixed-point definition code
huffman
parents: 33774
diff changeset
   190
7a1518c42c56 cleaned up; factored out fixed-point definition code
huffman
parents: 33774
diff changeset
   191
    (* project components of fixpoint *)
7a1518c42c56 cleaned up; factored out fixed-point definition code
huffman
parents: 33774
diff changeset
   192
    fun mk_projs (x::[]) t = [(x, t)]
7a1518c42c56 cleaned up; factored out fixed-point definition code
huffman
parents: 33774
diff changeset
   193
      | mk_projs (x::xs) t = (x, mk_fst t) :: mk_projs xs (mk_snd t);
7a1518c42c56 cleaned up; factored out fixed-point definition code
huffman
parents: 33774
diff changeset
   194
    val projs = mk_projs lhss fixpoint;
7a1518c42c56 cleaned up; factored out fixed-point definition code
huffman
parents: 33774
diff changeset
   195
7a1518c42c56 cleaned up; factored out fixed-point definition code
huffman
parents: 33774
diff changeset
   196
    (* convert parameters to lambda abstractions *)
7a1518c42c56 cleaned up; factored out fixed-point definition code
huffman
parents: 33774
diff changeset
   197
    fun mk_eqn (lhs, rhs) =
7a1518c42c56 cleaned up; factored out fixed-point definition code
huffman
parents: 33774
diff changeset
   198
        case lhs of
7a1518c42c56 cleaned up; factored out fixed-point definition code
huffman
parents: 33774
diff changeset
   199
          Const (@{const_name Rep_CFun}, _) $ f $ (x as Free _) =>
7a1518c42c56 cleaned up; factored out fixed-point definition code
huffman
parents: 33774
diff changeset
   200
            mk_eqn (f, big_lambda x rhs)
7a1518c42c56 cleaned up; factored out fixed-point definition code
huffman
parents: 33774
diff changeset
   201
        | Const _ => Logic.mk_equals (lhs, rhs)
7a1518c42c56 cleaned up; factored out fixed-point definition code
huffman
parents: 33774
diff changeset
   202
        | _ => raise TERM ("lhs not of correct form", [lhs, rhs]);
7a1518c42c56 cleaned up; factored out fixed-point definition code
huffman
parents: 33774
diff changeset
   203
    val eqns = map mk_eqn projs;
7a1518c42c56 cleaned up; factored out fixed-point definition code
huffman
parents: 33774
diff changeset
   204
7a1518c42c56 cleaned up; factored out fixed-point definition code
huffman
parents: 33774
diff changeset
   205
    (* register constant definitions *)
33777
69eae9bca167 get rid of numbers on thy variables
huffman
parents: 33776
diff changeset
   206
    val (fixdef_thms, thy) =
33775
7a1518c42c56 cleaned up; factored out fixed-point definition code
huffman
parents: 33774
diff changeset
   207
      (PureThy.add_defs false o map Thm.no_attributes)
7a1518c42c56 cleaned up; factored out fixed-point definition code
huffman
parents: 33774
diff changeset
   208
        (map (Binding.suffix_name "_def") binds ~~ eqns) thy;
7a1518c42c56 cleaned up; factored out fixed-point definition code
huffman
parents: 33774
diff changeset
   209
7a1518c42c56 cleaned up; factored out fixed-point definition code
huffman
parents: 33774
diff changeset
   210
    (* prove applied version of definitions *)
7a1518c42c56 cleaned up; factored out fixed-point definition code
huffman
parents: 33774
diff changeset
   211
    fun prove_proj (lhs, rhs) =
7a1518c42c56 cleaned up; factored out fixed-point definition code
huffman
parents: 33774
diff changeset
   212
      let
7a1518c42c56 cleaned up; factored out fixed-point definition code
huffman
parents: 33774
diff changeset
   213
        val tac = rewrite_goals_tac fixdef_thms THEN beta_tac 1;
7a1518c42c56 cleaned up; factored out fixed-point definition code
huffman
parents: 33774
diff changeset
   214
        val goal = Logic.mk_equals (lhs, rhs);
33777
69eae9bca167 get rid of numbers on thy variables
huffman
parents: 33776
diff changeset
   215
      in Goal.prove_global thy [] [] goal (K tac) end;
33775
7a1518c42c56 cleaned up; factored out fixed-point definition code
huffman
parents: 33774
diff changeset
   216
    val proj_thms = map prove_proj projs;
7a1518c42c56 cleaned up; factored out fixed-point definition code
huffman
parents: 33774
diff changeset
   217
7a1518c42c56 cleaned up; factored out fixed-point definition code
huffman
parents: 33774
diff changeset
   218
    (* mk_tuple lhss == fixpoint *)
7a1518c42c56 cleaned up; factored out fixed-point definition code
huffman
parents: 33774
diff changeset
   219
    fun pair_equalI (thm1, thm2) = @{thm Pair_equalI} OF [thm1, thm2];
7a1518c42c56 cleaned up; factored out fixed-point definition code
huffman
parents: 33774
diff changeset
   220
    val tuple_fixdef_thm = foldr1 pair_equalI proj_thms;
7a1518c42c56 cleaned up; factored out fixed-point definition code
huffman
parents: 33774
diff changeset
   221
7a1518c42c56 cleaned up; factored out fixed-point definition code
huffman
parents: 33774
diff changeset
   222
    val cont_thm =
33777
69eae9bca167 get rid of numbers on thy variables
huffman
parents: 33776
diff changeset
   223
      Goal.prove_global thy [] [] (mk_trp (mk_cont functional))
33775
7a1518c42c56 cleaned up; factored out fixed-point definition code
huffman
parents: 33774
diff changeset
   224
        (K (beta_tac 1));
7a1518c42c56 cleaned up; factored out fixed-point definition code
huffman
parents: 33774
diff changeset
   225
    val tuple_unfold_thm =
7a1518c42c56 cleaned up; factored out fixed-point definition code
huffman
parents: 33774
diff changeset
   226
      (@{thm def_cont_fix_eq} OF [tuple_fixdef_thm, cont_thm])
33777
69eae9bca167 get rid of numbers on thy variables
huffman
parents: 33776
diff changeset
   227
      |> LocalDefs.unfold (ProofContext.init thy) @{thms split_conv};
33775
7a1518c42c56 cleaned up; factored out fixed-point definition code
huffman
parents: 33774
diff changeset
   228
7a1518c42c56 cleaned up; factored out fixed-point definition code
huffman
parents: 33774
diff changeset
   229
    fun mk_unfold_thms [] thm = []
7a1518c42c56 cleaned up; factored out fixed-point definition code
huffman
parents: 33774
diff changeset
   230
      | mk_unfold_thms (n::[]) thm = [(n, thm)]
7a1518c42c56 cleaned up; factored out fixed-point definition code
huffman
parents: 33774
diff changeset
   231
      | mk_unfold_thms (n::ns) thm = let
7a1518c42c56 cleaned up; factored out fixed-point definition code
huffman
parents: 33774
diff changeset
   232
          val thmL = thm RS @{thm Pair_eqD1};
7a1518c42c56 cleaned up; factored out fixed-point definition code
huffman
parents: 33774
diff changeset
   233
          val thmR = thm RS @{thm Pair_eqD2};
7a1518c42c56 cleaned up; factored out fixed-point definition code
huffman
parents: 33774
diff changeset
   234
        in (n, thmL) :: mk_unfold_thms ns thmR end;
7a1518c42c56 cleaned up; factored out fixed-point definition code
huffman
parents: 33774
diff changeset
   235
    val unfold_binds = map (Binding.suffix_name "_unfold") binds;
7a1518c42c56 cleaned up; factored out fixed-point definition code
huffman
parents: 33774
diff changeset
   236
7a1518c42c56 cleaned up; factored out fixed-point definition code
huffman
parents: 33774
diff changeset
   237
    (* register unfold theorems *)
33777
69eae9bca167 get rid of numbers on thy variables
huffman
parents: 33776
diff changeset
   238
    val (unfold_thms, thy) =
33775
7a1518c42c56 cleaned up; factored out fixed-point definition code
huffman
parents: 33774
diff changeset
   239
      (PureThy.add_thms o map (Thm.no_attributes o apsnd Drule.standard))
33777
69eae9bca167 get rid of numbers on thy variables
huffman
parents: 33776
diff changeset
   240
        (mk_unfold_thms unfold_binds tuple_unfold_thm) thy;
33775
7a1518c42c56 cleaned up; factored out fixed-point definition code
huffman
parents: 33774
diff changeset
   241
  in
33788
481bc899febf automate isodefl proof
huffman
parents: 33785
diff changeset
   242
    ((proj_thms, unfold_thms), thy)
33775
7a1518c42c56 cleaned up; factored out fixed-point definition code
huffman
parents: 33774
diff changeset
   243
  end;
7a1518c42c56 cleaned up; factored out fixed-point definition code
huffman
parents: 33774
diff changeset
   244
7a1518c42c56 cleaned up; factored out fixed-point definition code
huffman
parents: 33774
diff changeset
   245
33774
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   246
(******************************************************************************)
33790
b2ff505e30f8 replace defl_tab and map_tab with theory data
huffman
parents: 33789
diff changeset
   247
(****************** deflation combinators and map functions *******************)
b2ff505e30f8 replace defl_tab and map_tab with theory data
huffman
parents: 33789
diff changeset
   248
(******************************************************************************)
33774
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   249
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   250
fun defl_of_typ
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   251
    (tab : term Symtab.table)
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   252
    (T : typ) : term =
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   253
  let
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   254
    fun is_closed_typ (Type (_, Ts)) = forall is_closed_typ Ts
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   255
      | is_closed_typ _ = false;
33785
2f2d9eb37084 automate definition of map functions; remove unused code
huffman
parents: 33784
diff changeset
   256
    fun defl_of (TFree (a, _)) = Free (Library.unprefix "'" a, deflT)
33774
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   257
      | defl_of (TVar _) = error ("defl_of_typ: TVar")
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   258
      | defl_of (T as Type (c, Ts)) =
33775
7a1518c42c56 cleaned up; factored out fixed-point definition code
huffman
parents: 33774
diff changeset
   259
        case Symtab.lookup tab c of
33774
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   260
          SOME t => Library.foldl mk_capply (t, map defl_of Ts)
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   261
        | NONE => if is_closed_typ T
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   262
                  then mk_Rep_of T
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   263
                  else error ("defl_of_typ: type variable under unsupported type constructor " ^ c);
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   264
  in defl_of T end;
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   265
33785
2f2d9eb37084 automate definition of map functions; remove unused code
huffman
parents: 33784
diff changeset
   266
fun map_of_typ
2f2d9eb37084 automate definition of map functions; remove unused code
huffman
parents: 33784
diff changeset
   267
    (tab : string Symtab.table)
2f2d9eb37084 automate definition of map functions; remove unused code
huffman
parents: 33784
diff changeset
   268
    (T : typ) : term =
33774
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   269
  let
33785
2f2d9eb37084 automate definition of map functions; remove unused code
huffman
parents: 33784
diff changeset
   270
    fun is_closed_typ (Type (_, Ts)) = forall is_closed_typ Ts
2f2d9eb37084 automate definition of map functions; remove unused code
huffman
parents: 33784
diff changeset
   271
      | is_closed_typ _ = false;
2f2d9eb37084 automate definition of map functions; remove unused code
huffman
parents: 33784
diff changeset
   272
    fun map_of (T as TFree (a, _)) = Free (Library.unprefix "'" a, T ->> T)
2f2d9eb37084 automate definition of map functions; remove unused code
huffman
parents: 33784
diff changeset
   273
      | map_of (T as TVar _) = error ("map_of_typ: TVar")
2f2d9eb37084 automate definition of map functions; remove unused code
huffman
parents: 33784
diff changeset
   274
      | map_of (T as Type (c, Ts)) =
2f2d9eb37084 automate definition of map functions; remove unused code
huffman
parents: 33784
diff changeset
   275
        case Symtab.lookup tab c of
2f2d9eb37084 automate definition of map functions; remove unused code
huffman
parents: 33784
diff changeset
   276
          SOME t => Library.foldl mk_capply (Const (t, mapT T), map map_of Ts)
2f2d9eb37084 automate definition of map functions; remove unused code
huffman
parents: 33784
diff changeset
   277
        | NONE => if is_closed_typ T
2f2d9eb37084 automate definition of map functions; remove unused code
huffman
parents: 33784
diff changeset
   278
                  then ID_const T
2f2d9eb37084 automate definition of map functions; remove unused code
huffman
parents: 33784
diff changeset
   279
                  else error ("map_of_typ: type variable under unsupported type constructor " ^ c);
2f2d9eb37084 automate definition of map functions; remove unused code
huffman
parents: 33784
diff changeset
   280
  in map_of T end;
2f2d9eb37084 automate definition of map functions; remove unused code
huffman
parents: 33784
diff changeset
   281
33774
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   282
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   283
(******************************************************************************)
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   284
(* prepare datatype specifications *)
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   285
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   286
fun read_typ thy str sorts =
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   287
  let
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   288
    val ctxt = ProofContext.init thy
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   289
      |> fold (Variable.declare_typ o TFree) sorts;
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   290
    val T = Syntax.read_typ ctxt str;
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   291
  in (T, Term.add_tfreesT T sorts) end;
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   292
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   293
fun cert_typ sign raw_T sorts =
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   294
  let
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   295
    val T = Type.no_tvars (Sign.certify_typ sign raw_T)
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   296
      handle TYPE (msg, _, _) => error msg;
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   297
    val sorts' = Term.add_tfreesT T sorts;
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   298
    val _ =
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   299
      case duplicates (op =) (map fst sorts') of
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   300
        [] => ()
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   301
      | dups => error ("Inconsistent sort constraints for " ^ commas dups)
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   302
  in (T, sorts') end;
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   303
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   304
fun gen_domain_isomorphism
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   305
    (prep_typ: theory -> 'a -> (string * sort) list -> typ * (string * sort) list)
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   306
    (doms_raw: (string list * binding * mixfix * 'a) list)
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   307
    (thy: theory)
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   308
    : theory =
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   309
  let
33792
002e0e017311 change Theory.requires
huffman
parents: 33791
diff changeset
   310
    val _ = Theory.requires thy "Representable" "domain isomorphisms";
33774
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   311
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   312
    (* this theory is used just for parsing *)
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   313
    val tmp_thy = thy |>
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   314
      Theory.copy |>
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   315
      Sign.add_types (map (fn (tvs, tname, mx, _) =>
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   316
        (tname, length tvs, mx)) doms_raw);
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   317
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   318
    fun prep_dom thy (vs, t, mx, typ_raw) sorts =
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   319
      let val (typ, sorts') = prep_typ thy typ_raw sorts
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   320
      in ((vs, t, mx, typ), sorts') end;
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   321
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   322
    val (doms : (string list * binding * mixfix * typ) list,
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   323
         sorts : (string * sort) list) =
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   324
      fold_map (prep_dom tmp_thy) doms_raw [];
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   325
33775
7a1518c42c56 cleaned up; factored out fixed-point definition code
huffman
parents: 33774
diff changeset
   326
    (* domain equations *)
7a1518c42c56 cleaned up; factored out fixed-point definition code
huffman
parents: 33774
diff changeset
   327
    fun mk_dom_eqn (vs, tbind, mx, rhs) =
7a1518c42c56 cleaned up; factored out fixed-point definition code
huffman
parents: 33774
diff changeset
   328
      let fun arg v = TFree (v, the (AList.lookup (op =) sorts v));
7a1518c42c56 cleaned up; factored out fixed-point definition code
huffman
parents: 33774
diff changeset
   329
      in (Type (Sign.full_name tmp_thy tbind, map arg vs), rhs) end;
7a1518c42c56 cleaned up; factored out fixed-point definition code
huffman
parents: 33774
diff changeset
   330
    val dom_eqns = map mk_dom_eqn doms;
7a1518c42c56 cleaned up; factored out fixed-point definition code
huffman
parents: 33774
diff changeset
   331
7a1518c42c56 cleaned up; factored out fixed-point definition code
huffman
parents: 33774
diff changeset
   332
    (* check for valid type parameters *)
33774
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   333
    val (tyvars, _, _, _)::_ = doms;
33775
7a1518c42c56 cleaned up; factored out fixed-point definition code
huffman
parents: 33774
diff changeset
   334
    val new_doms = map (fn (tvs, tname, mx, _) =>
33774
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   335
      let val full_tname = Sign.full_name tmp_thy tname
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   336
      in
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   337
        (case duplicates (op =) tvs of
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   338
          [] =>
33775
7a1518c42c56 cleaned up; factored out fixed-point definition code
huffman
parents: 33774
diff changeset
   339
            if eq_set (op =) (tyvars, tvs) then (full_tname, tvs)
33774
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   340
            else error ("Mutually recursive domains must have same type parameters")
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   341
        | dups => error ("Duplicate parameter(s) for domain " ^ quote (Binding.str_of tname) ^
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   342
            " : " ^ commas dups))
33775
7a1518c42c56 cleaned up; factored out fixed-point definition code
huffman
parents: 33774
diff changeset
   343
      end) doms;
33782
cdb3ca1a765d prove isomorphism and isodefl rules
huffman
parents: 33778
diff changeset
   344
    val dom_binds = map (fn (_, tbind, _, _) => tbind) doms;
33774
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   345
33784
7e434813752f change naming convention for deflation combinators
huffman
parents: 33782
diff changeset
   346
    (* declare deflation combinator constants *)
7e434813752f change naming convention for deflation combinators
huffman
parents: 33782
diff changeset
   347
    fun declare_defl_const (vs, tbind, mx, rhs) thy =
33774
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   348
      let
33784
7e434813752f change naming convention for deflation combinators
huffman
parents: 33782
diff changeset
   349
        val defl_type = Library.foldr cfunT (map (K deflT) vs, deflT);
7e434813752f change naming convention for deflation combinators
huffman
parents: 33782
diff changeset
   350
        val defl_bind = Binding.suffix_name "_defl" tbind;
33774
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   351
      in
33784
7e434813752f change naming convention for deflation combinators
huffman
parents: 33782
diff changeset
   352
        Sign.declare_const ((defl_bind, defl_type), NoSyn) thy
33774
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   353
      end;
33784
7e434813752f change naming convention for deflation combinators
huffman
parents: 33782
diff changeset
   354
    val (defl_consts, thy) = fold_map declare_defl_const doms thy;
33774
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   355
33775
7a1518c42c56 cleaned up; factored out fixed-point definition code
huffman
parents: 33774
diff changeset
   356
    (* defining equations for type combinators *)
33790
b2ff505e30f8 replace defl_tab and map_tab with theory data
huffman
parents: 33789
diff changeset
   357
    val defl_tab1 = DeflData.get thy;
33775
7a1518c42c56 cleaned up; factored out fixed-point definition code
huffman
parents: 33774
diff changeset
   358
    val defl_tab2 =
33784
7e434813752f change naming convention for deflation combinators
huffman
parents: 33782
diff changeset
   359
      Symtab.make (map (fst o dest_Type o fst) dom_eqns ~~ defl_consts);
33775
7a1518c42c56 cleaned up; factored out fixed-point definition code
huffman
parents: 33774
diff changeset
   360
    val defl_tab' = Symtab.merge (K true) (defl_tab1, defl_tab2);
33790
b2ff505e30f8 replace defl_tab and map_tab with theory data
huffman
parents: 33789
diff changeset
   361
    val thy = DeflData.put defl_tab' thy;
33776
5048b02c2bbb automate proofs of REP equations
huffman
parents: 33775
diff changeset
   362
    fun mk_defl_spec (lhsT, rhsT) =
33785
2f2d9eb37084 automate definition of map functions; remove unused code
huffman
parents: 33784
diff changeset
   363
      mk_eqs (defl_of_typ defl_tab' lhsT,
2f2d9eb37084 automate definition of map functions; remove unused code
huffman
parents: 33784
diff changeset
   364
              defl_of_typ defl_tab' rhsT);
33775
7a1518c42c56 cleaned up; factored out fixed-point definition code
huffman
parents: 33774
diff changeset
   365
    val defl_specs = map mk_defl_spec dom_eqns;
33774
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   366
33784
7e434813752f change naming convention for deflation combinators
huffman
parents: 33782
diff changeset
   367
    (* register recursive definition of deflation combinators *)
7e434813752f change naming convention for deflation combinators
huffman
parents: 33782
diff changeset
   368
    val defl_binds = map (Binding.suffix_name "_defl") dom_binds;
33788
481bc899febf automate isodefl proof
huffman
parents: 33785
diff changeset
   369
    val ((defl_apply_thms, defl_unfold_thms), thy) =
481bc899febf automate isodefl proof
huffman
parents: 33785
diff changeset
   370
      add_fixdefs (defl_binds ~~ defl_specs) thy;
33774
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   371
33776
5048b02c2bbb automate proofs of REP equations
huffman
parents: 33775
diff changeset
   372
    (* define types using deflation combinators *)
33784
7e434813752f change naming convention for deflation combinators
huffman
parents: 33782
diff changeset
   373
    fun make_repdef ((vs, tbind, mx, _), defl_const) thy =
33774
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   374
      let
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   375
        fun tfree a = TFree (a, the (AList.lookup (op =) sorts a))
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   376
        val reps = map (mk_Rep_of o tfree) vs;
33784
7e434813752f change naming convention for deflation combinators
huffman
parents: 33782
diff changeset
   377
        val defl = Library.foldl mk_capply (defl_const, reps);
33777
69eae9bca167 get rid of numbers on thy variables
huffman
parents: 33776
diff changeset
   378
        val ((_, _, _, {REP, ...}), thy) =
33774
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   379
          Repdef.add_repdef false NONE (tbind, vs, mx) defl NONE thy;
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   380
      in
33777
69eae9bca167 get rid of numbers on thy variables
huffman
parents: 33776
diff changeset
   381
        (REP, thy)
33774
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   382
      end;
33784
7e434813752f change naming convention for deflation combinators
huffman
parents: 33782
diff changeset
   383
    val (REP_thms, thy) = fold_map make_repdef (doms ~~ defl_consts) thy;
33791
fef59343b4b3 use theory data for REP_simps and isodefl_rules
huffman
parents: 33790
diff changeset
   384
    val thy = RepData.map (fold Thm.add_thm REP_thms) thy;
33776
5048b02c2bbb automate proofs of REP equations
huffman
parents: 33775
diff changeset
   385
5048b02c2bbb automate proofs of REP equations
huffman
parents: 33775
diff changeset
   386
    (* prove REP equations *)
33782
cdb3ca1a765d prove isomorphism and isodefl rules
huffman
parents: 33778
diff changeset
   387
    fun mk_REP_eq_thm (lhsT, rhsT) =
33776
5048b02c2bbb automate proofs of REP equations
huffman
parents: 33775
diff changeset
   388
      let
5048b02c2bbb automate proofs of REP equations
huffman
parents: 33775
diff changeset
   389
        val goal = mk_eqs (mk_Rep_of lhsT, mk_Rep_of rhsT);
33791
fef59343b4b3 use theory data for REP_simps and isodefl_rules
huffman
parents: 33790
diff changeset
   390
        val REP_simps = RepData.get thy;
33776
5048b02c2bbb automate proofs of REP equations
huffman
parents: 33775
diff changeset
   391
        val tac =
5048b02c2bbb automate proofs of REP equations
huffman
parents: 33775
diff changeset
   392
          simp_tac (HOL_basic_ss addsimps REP_simps) 1
33784
7e434813752f change naming convention for deflation combinators
huffman
parents: 33782
diff changeset
   393
          THEN resolve_tac defl_unfold_thms 1;
33776
5048b02c2bbb automate proofs of REP equations
huffman
parents: 33775
diff changeset
   394
      in
33777
69eae9bca167 get rid of numbers on thy variables
huffman
parents: 33776
diff changeset
   395
        Goal.prove_global thy [] [] goal (K tac)
33776
5048b02c2bbb automate proofs of REP equations
huffman
parents: 33775
diff changeset
   396
      end;
33782
cdb3ca1a765d prove isomorphism and isodefl rules
huffman
parents: 33778
diff changeset
   397
    val REP_eq_thms = map mk_REP_eq_thm dom_eqns;
cdb3ca1a765d prove isomorphism and isodefl rules
huffman
parents: 33778
diff changeset
   398
cdb3ca1a765d prove isomorphism and isodefl rules
huffman
parents: 33778
diff changeset
   399
    (* register REP equations *)
cdb3ca1a765d prove isomorphism and isodefl rules
huffman
parents: 33778
diff changeset
   400
    val REP_eq_binds = map (Binding.prefix_name "REP_eq_") dom_binds;
cdb3ca1a765d prove isomorphism and isodefl rules
huffman
parents: 33778
diff changeset
   401
    val (_, thy) = thy |>
cdb3ca1a765d prove isomorphism and isodefl rules
huffman
parents: 33778
diff changeset
   402
      (PureThy.add_thms o map Thm.no_attributes)
cdb3ca1a765d prove isomorphism and isodefl rules
huffman
parents: 33778
diff changeset
   403
        (REP_eq_binds ~~ REP_eq_thms);
33776
5048b02c2bbb automate proofs of REP equations
huffman
parents: 33775
diff changeset
   404
33778
9121ea165a40 automate definition of rep/abs functions
huffman
parents: 33777
diff changeset
   405
    (* define rep/abs functions *)
33782
cdb3ca1a765d prove isomorphism and isodefl rules
huffman
parents: 33778
diff changeset
   406
    fun mk_rep_abs (tbind, (lhsT, rhsT)) thy =
33778
9121ea165a40 automate definition of rep/abs functions
huffman
parents: 33777
diff changeset
   407
      let
9121ea165a40 automate definition of rep/abs functions
huffman
parents: 33777
diff changeset
   408
        val rep_type = cfunT (lhsT, rhsT);
33782
cdb3ca1a765d prove isomorphism and isodefl rules
huffman
parents: 33778
diff changeset
   409
        val abs_type = cfunT (rhsT, lhsT);
33778
9121ea165a40 automate definition of rep/abs functions
huffman
parents: 33777
diff changeset
   410
        val rep_bind = Binding.suffix_name "_rep" tbind;
9121ea165a40 automate definition of rep/abs functions
huffman
parents: 33777
diff changeset
   411
        val abs_bind = Binding.suffix_name "_abs" tbind;
9121ea165a40 automate definition of rep/abs functions
huffman
parents: 33777
diff changeset
   412
        val (rep_const, thy) = thy |>
9121ea165a40 automate definition of rep/abs functions
huffman
parents: 33777
diff changeset
   413
          Sign.declare_const ((rep_bind, rep_type), NoSyn);
9121ea165a40 automate definition of rep/abs functions
huffman
parents: 33777
diff changeset
   414
        val (abs_const, thy) = thy |>
9121ea165a40 automate definition of rep/abs functions
huffman
parents: 33777
diff changeset
   415
          Sign.declare_const ((abs_bind, abs_type), NoSyn);
9121ea165a40 automate definition of rep/abs functions
huffman
parents: 33777
diff changeset
   416
        val rep_eqn = Logic.mk_equals (rep_const, coerce_const rep_type);
9121ea165a40 automate definition of rep/abs functions
huffman
parents: 33777
diff changeset
   417
        val abs_eqn = Logic.mk_equals (abs_const, coerce_const abs_type);
9121ea165a40 automate definition of rep/abs functions
huffman
parents: 33777
diff changeset
   418
        val ([rep_def, abs_def], thy) = thy |>
9121ea165a40 automate definition of rep/abs functions
huffman
parents: 33777
diff changeset
   419
          (PureThy.add_defs false o map Thm.no_attributes)
9121ea165a40 automate definition of rep/abs functions
huffman
parents: 33777
diff changeset
   420
            [(Binding.suffix_name "_rep_def" tbind, rep_eqn),
9121ea165a40 automate definition of rep/abs functions
huffman
parents: 33777
diff changeset
   421
             (Binding.suffix_name "_abs_def" tbind, abs_eqn)];
9121ea165a40 automate definition of rep/abs functions
huffman
parents: 33777
diff changeset
   422
      in
33785
2f2d9eb37084 automate definition of map functions; remove unused code
huffman
parents: 33784
diff changeset
   423
        (((rep_const, abs_const), (rep_def, abs_def)), thy)
33778
9121ea165a40 automate definition of rep/abs functions
huffman
parents: 33777
diff changeset
   424
      end;
33785
2f2d9eb37084 automate definition of map functions; remove unused code
huffman
parents: 33784
diff changeset
   425
    val ((rep_abs_consts, rep_abs_defs), thy) = thy
2f2d9eb37084 automate definition of map functions; remove unused code
huffman
parents: 33784
diff changeset
   426
      |> fold_map mk_rep_abs (dom_binds ~~ dom_eqns)
2f2d9eb37084 automate definition of map functions; remove unused code
huffman
parents: 33784
diff changeset
   427
      |>> ListPair.unzip;
33782
cdb3ca1a765d prove isomorphism and isodefl rules
huffman
parents: 33778
diff changeset
   428
cdb3ca1a765d prove isomorphism and isodefl rules
huffman
parents: 33778
diff changeset
   429
    (* prove isomorphism and isodefl rules *)
cdb3ca1a765d prove isomorphism and isodefl rules
huffman
parents: 33778
diff changeset
   430
    fun mk_iso_thms ((tbind, REP_eq), (rep_def, abs_def)) thy =
cdb3ca1a765d prove isomorphism and isodefl rules
huffman
parents: 33778
diff changeset
   431
      let
cdb3ca1a765d prove isomorphism and isodefl rules
huffman
parents: 33778
diff changeset
   432
        fun make thm = Drule.standard (thm OF [REP_eq, abs_def, rep_def]);
cdb3ca1a765d prove isomorphism and isodefl rules
huffman
parents: 33778
diff changeset
   433
        val rep_iso_thm = make @{thm domain_rep_iso};
cdb3ca1a765d prove isomorphism and isodefl rules
huffman
parents: 33778
diff changeset
   434
        val abs_iso_thm = make @{thm domain_abs_iso};
cdb3ca1a765d prove isomorphism and isodefl rules
huffman
parents: 33778
diff changeset
   435
        val isodefl_thm = make @{thm isodefl_abs_rep};
cdb3ca1a765d prove isomorphism and isodefl rules
huffman
parents: 33778
diff changeset
   436
        val rep_iso_bind = Binding.suffix_name "_rep_iso" tbind;
cdb3ca1a765d prove isomorphism and isodefl rules
huffman
parents: 33778
diff changeset
   437
        val abs_iso_bind = Binding.suffix_name "_abs_iso" tbind;
cdb3ca1a765d prove isomorphism and isodefl rules
huffman
parents: 33778
diff changeset
   438
        val isodefl_bind = Binding.prefix_name "isodefl_abs_rep_" tbind;
cdb3ca1a765d prove isomorphism and isodefl rules
huffman
parents: 33778
diff changeset
   439
        val (_, thy) = thy |>
cdb3ca1a765d prove isomorphism and isodefl rules
huffman
parents: 33778
diff changeset
   440
          (PureThy.add_thms o map Thm.no_attributes)
cdb3ca1a765d prove isomorphism and isodefl rules
huffman
parents: 33778
diff changeset
   441
            [(rep_iso_bind, rep_iso_thm),
cdb3ca1a765d prove isomorphism and isodefl rules
huffman
parents: 33778
diff changeset
   442
             (abs_iso_bind, abs_iso_thm),
cdb3ca1a765d prove isomorphism and isodefl rules
huffman
parents: 33778
diff changeset
   443
             (isodefl_bind, isodefl_thm)];
cdb3ca1a765d prove isomorphism and isodefl rules
huffman
parents: 33778
diff changeset
   444
      in
cdb3ca1a765d prove isomorphism and isodefl rules
huffman
parents: 33778
diff changeset
   445
        (((rep_iso_thm, abs_iso_thm), isodefl_thm), thy)
cdb3ca1a765d prove isomorphism and isodefl rules
huffman
parents: 33778
diff changeset
   446
      end;
cdb3ca1a765d prove isomorphism and isodefl rules
huffman
parents: 33778
diff changeset
   447
    val ((iso_thms, isodefl_abs_rep_thms), thy) = thy
cdb3ca1a765d prove isomorphism and isodefl rules
huffman
parents: 33778
diff changeset
   448
      |> fold_map mk_iso_thms (dom_binds ~~ REP_eq_thms ~~ rep_abs_defs)
cdb3ca1a765d prove isomorphism and isodefl rules
huffman
parents: 33778
diff changeset
   449
      |>> ListPair.unzip;
33778
9121ea165a40 automate definition of rep/abs functions
huffman
parents: 33777
diff changeset
   450
33785
2f2d9eb37084 automate definition of map functions; remove unused code
huffman
parents: 33784
diff changeset
   451
    (* declare map functions *)
2f2d9eb37084 automate definition of map functions; remove unused code
huffman
parents: 33784
diff changeset
   452
    fun declare_map_const (tbind, (lhsT, rhsT)) thy =
2f2d9eb37084 automate definition of map functions; remove unused code
huffman
parents: 33784
diff changeset
   453
      let
2f2d9eb37084 automate definition of map functions; remove unused code
huffman
parents: 33784
diff changeset
   454
        val map_type = mapT lhsT;
2f2d9eb37084 automate definition of map functions; remove unused code
huffman
parents: 33784
diff changeset
   455
        val map_bind = Binding.suffix_name "_map" tbind;
2f2d9eb37084 automate definition of map functions; remove unused code
huffman
parents: 33784
diff changeset
   456
      in
2f2d9eb37084 automate definition of map functions; remove unused code
huffman
parents: 33784
diff changeset
   457
        Sign.declare_const ((map_bind, map_type), NoSyn) thy
2f2d9eb37084 automate definition of map functions; remove unused code
huffman
parents: 33784
diff changeset
   458
      end;
2f2d9eb37084 automate definition of map functions; remove unused code
huffman
parents: 33784
diff changeset
   459
    val (map_consts, thy) = thy |>
2f2d9eb37084 automate definition of map functions; remove unused code
huffman
parents: 33784
diff changeset
   460
      fold_map declare_map_const (dom_binds ~~ dom_eqns);
2f2d9eb37084 automate definition of map functions; remove unused code
huffman
parents: 33784
diff changeset
   461
2f2d9eb37084 automate definition of map functions; remove unused code
huffman
parents: 33784
diff changeset
   462
    (* defining equations for map functions *)
33790
b2ff505e30f8 replace defl_tab and map_tab with theory data
huffman
parents: 33789
diff changeset
   463
    val map_tab1 = MapData.get thy;
33785
2f2d9eb37084 automate definition of map functions; remove unused code
huffman
parents: 33784
diff changeset
   464
    val map_tab2 =
2f2d9eb37084 automate definition of map functions; remove unused code
huffman
parents: 33784
diff changeset
   465
      Symtab.make (map (fst o dest_Type o fst) dom_eqns
2f2d9eb37084 automate definition of map functions; remove unused code
huffman
parents: 33784
diff changeset
   466
                   ~~ map (fst o dest_Const) map_consts);
2f2d9eb37084 automate definition of map functions; remove unused code
huffman
parents: 33784
diff changeset
   467
    val map_tab' = Symtab.merge (K true) (map_tab1, map_tab2);
33790
b2ff505e30f8 replace defl_tab and map_tab with theory data
huffman
parents: 33789
diff changeset
   468
    val thy = MapData.put map_tab' thy;
33785
2f2d9eb37084 automate definition of map functions; remove unused code
huffman
parents: 33784
diff changeset
   469
    fun mk_map_spec ((rep_const, abs_const), (lhsT, rhsT)) =
2f2d9eb37084 automate definition of map functions; remove unused code
huffman
parents: 33784
diff changeset
   470
      let
2f2d9eb37084 automate definition of map functions; remove unused code
huffman
parents: 33784
diff changeset
   471
        val lhs = map_of_typ map_tab' lhsT;
2f2d9eb37084 automate definition of map functions; remove unused code
huffman
parents: 33784
diff changeset
   472
        val body = map_of_typ map_tab' rhsT;
2f2d9eb37084 automate definition of map functions; remove unused code
huffman
parents: 33784
diff changeset
   473
        val rhs = mk_cfcomp (abs_const, mk_cfcomp (body, rep_const));
2f2d9eb37084 automate definition of map functions; remove unused code
huffman
parents: 33784
diff changeset
   474
      in mk_eqs (lhs, rhs) end;
2f2d9eb37084 automate definition of map functions; remove unused code
huffman
parents: 33784
diff changeset
   475
    val map_specs = map mk_map_spec (rep_abs_consts ~~ dom_eqns);
2f2d9eb37084 automate definition of map functions; remove unused code
huffman
parents: 33784
diff changeset
   476
2f2d9eb37084 automate definition of map functions; remove unused code
huffman
parents: 33784
diff changeset
   477
    (* register recursive definition of map functions *)
2f2d9eb37084 automate definition of map functions; remove unused code
huffman
parents: 33784
diff changeset
   478
    val map_binds = map (Binding.suffix_name "_map") dom_binds;
33788
481bc899febf automate isodefl proof
huffman
parents: 33785
diff changeset
   479
    val ((map_apply_thms, map_unfold_thms), thy) =
481bc899febf automate isodefl proof
huffman
parents: 33785
diff changeset
   480
      add_fixdefs (map_binds ~~ map_specs) thy;
481bc899febf automate isodefl proof
huffman
parents: 33785
diff changeset
   481
481bc899febf automate isodefl proof
huffman
parents: 33785
diff changeset
   482
    (* prove isodefl rules for map functions *)
481bc899febf automate isodefl proof
huffman
parents: 33785
diff changeset
   483
    val isodefl_thm =
481bc899febf automate isodefl proof
huffman
parents: 33785
diff changeset
   484
      let
481bc899febf automate isodefl proof
huffman
parents: 33785
diff changeset
   485
        fun unprime a = Library.unprefix "'" a;
481bc899febf automate isodefl proof
huffman
parents: 33785
diff changeset
   486
        fun mk_d (TFree (a, _)) = Free ("d" ^ unprime a, deflT);
481bc899febf automate isodefl proof
huffman
parents: 33785
diff changeset
   487
        fun mk_f (T as TFree (a, _)) = Free ("f" ^ unprime a, T ->> T);
481bc899febf automate isodefl proof
huffman
parents: 33785
diff changeset
   488
        fun mk_assm T = mk_trp (isodefl_const T $ mk_f T $ mk_d T);
481bc899febf automate isodefl proof
huffman
parents: 33785
diff changeset
   489
        fun mk_goal ((map_const, defl_const), (T as Type (c, Ts), rhsT)) =
481bc899febf automate isodefl proof
huffman
parents: 33785
diff changeset
   490
          let
481bc899febf automate isodefl proof
huffman
parents: 33785
diff changeset
   491
            val map_term = Library.foldl mk_capply (map_const, map mk_f Ts);
481bc899febf automate isodefl proof
huffman
parents: 33785
diff changeset
   492
            val defl_term = Library.foldl mk_capply (defl_const, map mk_d Ts);
481bc899febf automate isodefl proof
huffman
parents: 33785
diff changeset
   493
          in isodefl_const T $ map_term $ defl_term end;
481bc899febf automate isodefl proof
huffman
parents: 33785
diff changeset
   494
        val assms = (map mk_assm o snd o dest_Type o fst o hd) dom_eqns;
481bc899febf automate isodefl proof
huffman
parents: 33785
diff changeset
   495
        val goals = map mk_goal (map_consts ~~ defl_consts ~~ dom_eqns);
481bc899febf automate isodefl proof
huffman
parents: 33785
diff changeset
   496
        val goal = mk_trp (foldr1 HOLogic.mk_conj goals);
481bc899febf automate isodefl proof
huffman
parents: 33785
diff changeset
   497
        val start_thms =
481bc899febf automate isodefl proof
huffman
parents: 33785
diff changeset
   498
          @{thm split_def} :: defl_apply_thms @ map_apply_thms;
481bc899febf automate isodefl proof
huffman
parents: 33785
diff changeset
   499
        val adm_rules =
481bc899febf automate isodefl proof
huffman
parents: 33785
diff changeset
   500
          @{thms adm_conj adm_isodefl cont2cont_fst cont2cont_snd cont_id};
481bc899febf automate isodefl proof
huffman
parents: 33785
diff changeset
   501
        val bottom_rules =
481bc899febf automate isodefl proof
huffman
parents: 33785
diff changeset
   502
          @{thms fst_strict snd_strict isodefl_bottom simp_thms};
481bc899febf automate isodefl proof
huffman
parents: 33785
diff changeset
   503
        val isodefl_rules =
33791
fef59343b4b3 use theory data for REP_simps and isodefl_rules
huffman
parents: 33790
diff changeset
   504
          @{thms conjI isodefl_ID_REP}
fef59343b4b3 use theory data for REP_simps and isodefl_rules
huffman
parents: 33790
diff changeset
   505
          @ isodefl_abs_rep_thms
fef59343b4b3 use theory data for REP_simps and isodefl_rules
huffman
parents: 33790
diff changeset
   506
          @ IsodeflData.get thy;
33788
481bc899febf automate isodefl proof
huffman
parents: 33785
diff changeset
   507
        fun tacf {prems, ...} = EVERY
481bc899febf automate isodefl proof
huffman
parents: 33785
diff changeset
   508
          [simp_tac (HOL_basic_ss addsimps start_thms) 1,
33791
fef59343b4b3 use theory data for REP_simps and isodefl_rules
huffman
parents: 33790
diff changeset
   509
           (* FIXME: how reliable is unification here? *)
fef59343b4b3 use theory data for REP_simps and isodefl_rules
huffman
parents: 33790
diff changeset
   510
           (* Maybe I should instantiate the rule. *)
33788
481bc899febf automate isodefl proof
huffman
parents: 33785
diff changeset
   511
           rtac @{thm parallel_fix_ind} 1,
481bc899febf automate isodefl proof
huffman
parents: 33785
diff changeset
   512
           REPEAT (resolve_tac adm_rules 1),
481bc899febf automate isodefl proof
huffman
parents: 33785
diff changeset
   513
           simp_tac (HOL_basic_ss addsimps bottom_rules) 1,
481bc899febf automate isodefl proof
huffman
parents: 33785
diff changeset
   514
           simp_tac beta_ss 1,
481bc899febf automate isodefl proof
huffman
parents: 33785
diff changeset
   515
           simp_tac (HOL_basic_ss addsimps @{thms fst_conv snd_conv}) 1,
481bc899febf automate isodefl proof
huffman
parents: 33785
diff changeset
   516
           REPEAT (etac @{thm conjE} 1),
481bc899febf automate isodefl proof
huffman
parents: 33785
diff changeset
   517
           REPEAT (resolve_tac (isodefl_rules @ prems) 1 ORELSE atac 1)];
481bc899febf automate isodefl proof
huffman
parents: 33785
diff changeset
   518
      in
481bc899febf automate isodefl proof
huffman
parents: 33785
diff changeset
   519
        Goal.prove_global thy [] assms goal tacf
481bc899febf automate isodefl proof
huffman
parents: 33785
diff changeset
   520
      end;
33789
c3fbdff7aed0 separate conjuncts of isodefl theorem
huffman
parents: 33788
diff changeset
   521
    val isodefl_binds = map (Binding.prefix_name "isodefl_") dom_binds;
c3fbdff7aed0 separate conjuncts of isodefl theorem
huffman
parents: 33788
diff changeset
   522
    fun conjuncts [] thm = []
c3fbdff7aed0 separate conjuncts of isodefl theorem
huffman
parents: 33788
diff changeset
   523
      | conjuncts (n::[]) thm = [(n, thm)]
c3fbdff7aed0 separate conjuncts of isodefl theorem
huffman
parents: 33788
diff changeset
   524
      | conjuncts (n::ns) thm = let
c3fbdff7aed0 separate conjuncts of isodefl theorem
huffman
parents: 33788
diff changeset
   525
          val thmL = thm RS @{thm conjunct1};
c3fbdff7aed0 separate conjuncts of isodefl theorem
huffman
parents: 33788
diff changeset
   526
          val thmR = thm RS @{thm conjunct2};
c3fbdff7aed0 separate conjuncts of isodefl theorem
huffman
parents: 33788
diff changeset
   527
        in (n, thmL):: conjuncts ns thmR end;
c3fbdff7aed0 separate conjuncts of isodefl theorem
huffman
parents: 33788
diff changeset
   528
    val (isodefl_thms, thy) = thy |>
c3fbdff7aed0 separate conjuncts of isodefl theorem
huffman
parents: 33788
diff changeset
   529
      (PureThy.add_thms o map (Thm.no_attributes o apsnd Drule.standard))
c3fbdff7aed0 separate conjuncts of isodefl theorem
huffman
parents: 33788
diff changeset
   530
        (conjuncts isodefl_binds isodefl_thm);
33791
fef59343b4b3 use theory data for REP_simps and isodefl_rules
huffman
parents: 33790
diff changeset
   531
    val thy = IsodeflData.map (fold Thm.add_thm isodefl_thms) thy;
33774
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   532
  in
33777
69eae9bca167 get rid of numbers on thy variables
huffman
parents: 33776
diff changeset
   533
    thy
33774
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   534
  end;
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   535
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   536
val domain_isomorphism = gen_domain_isomorphism cert_typ;
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   537
val domain_isomorphism_cmd = gen_domain_isomorphism read_typ;
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   538
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   539
(******************************************************************************)
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   540
(******************************** outer syntax ********************************)
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   541
(******************************************************************************)
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   542
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   543
local
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   544
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   545
structure P = OuterParse and K = OuterKeyword
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   546
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   547
val parse_domain_iso : (string list * binding * mixfix * string) parser =
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   548
  (P.type_args -- P.binding -- P.opt_infix -- (P.$$$ "=" |-- P.typ))
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   549
    >> (fn (((vs, t), mx), rhs) => (vs, t, mx, rhs));
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   550
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   551
val parse_domain_isos = P.and_list1 parse_domain_iso;
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   552
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   553
in
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   554
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   555
val _ =
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   556
  OuterSyntax.command "domain_isomorphism" "define domain isomorphisms (HOLCF)" K.thy_decl
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   557
    (parse_domain_isos >> (Toplevel.theory o domain_isomorphism_cmd));
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   558
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   559
end;
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   560
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   561
end;