src/HOLCF/Tools/Domain/domain_isomorphism.ML
author huffman
Wed, 18 Nov 2009 15:54:47 -0800
changeset 33777 69eae9bca167
parent 33776 5048b02c2bbb
child 33778 9121ea165a40
permissions -rw-r--r--
get rid of numbers on thy variables
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:
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
    10
     (string list * binding * mixfix * typ) list -> theory -> theory
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
    11
  val domain_isomorphism_cmd:
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
    12
     (string list * binding * mixfix * string) list -> theory -> theory
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
    13
end;
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
    14
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
    15
structure Domain_Isomorphism :> DOMAIN_ISOMORPHISM =
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
    16
struct
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
    17
33775
7a1518c42c56 cleaned up; factored out fixed-point definition code
huffman
parents: 33774
diff changeset
    18
val beta_ss =
7a1518c42c56 cleaned up; factored out fixed-point definition code
huffman
parents: 33774
diff changeset
    19
  HOL_basic_ss
7a1518c42c56 cleaned up; factored out fixed-point definition code
huffman
parents: 33774
diff changeset
    20
    addsimps simp_thms
7a1518c42c56 cleaned up; factored out fixed-point definition code
huffman
parents: 33774
diff changeset
    21
    addsimps [@{thm beta_cfun}]
7a1518c42c56 cleaned up; factored out fixed-point definition code
huffman
parents: 33774
diff changeset
    22
    addsimprocs [@{simproc cont_proc}];
7a1518c42c56 cleaned up; factored out fixed-point definition code
huffman
parents: 33774
diff changeset
    23
7a1518c42c56 cleaned up; factored out fixed-point definition code
huffman
parents: 33774
diff changeset
    24
val beta_tac = simp_tac beta_ss;
7a1518c42c56 cleaned up; factored out fixed-point definition code
huffman
parents: 33774
diff changeset
    25
33774
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
    26
(******************************************************************************)
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
    27
(******************************* building types *******************************)
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
    28
(******************************************************************************)
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
    29
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
    30
(* ->> is taken from holcf_logic.ML *)
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
    31
fun cfunT (T, U) = Type(@{type_name "->"}, [T, U]);
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
    32
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
    33
infixr 6 ->>; val (op ->>) = cfunT;
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
    34
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
    35
fun dest_cfunT (Type(@{type_name "->"}, [T, U])) = (T, U)
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
    36
  | dest_cfunT T = raise TYPE ("dest_cfunT", [T], []);
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
    37
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
    38
fun tupleT [] = HOLogic.unitT
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
    39
  | tupleT [T] = T
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
    40
  | tupleT (T :: Ts) = HOLogic.mk_prodT (T, tupleT Ts);
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
    41
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
    42
val deflT = @{typ "udom alg_defl"};
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
    43
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
    44
(******************************************************************************)
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
    45
(******************************* building terms *******************************)
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
    46
(******************************************************************************)
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
    47
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
    48
(* builds the expression (v1,v2,..,vn) *)
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
    49
fun mk_tuple [] = HOLogic.unit
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
    50
|   mk_tuple (t::[]) = t
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
    51
|   mk_tuple (t::ts) = HOLogic.mk_prod (t, mk_tuple ts);
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
    52
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
    53
(* builds the expression (%(v1,v2,..,vn). rhs) *)
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
    54
fun lambda_tuple [] rhs = Term.lambda (Free("unit", HOLogic.unitT)) rhs
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
    55
  | lambda_tuple (v::[]) rhs = Term.lambda v rhs
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
    56
  | lambda_tuple (v::vs) rhs =
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
    57
      HOLogic.mk_split (Term.lambda v (lambda_tuple vs rhs));
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
    58
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
    59
(* continuous application and abstraction *)
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
    60
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
    61
fun capply_const (S, T) =
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
    62
  Const(@{const_name Rep_CFun}, (S ->> T) --> (S --> T));
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
    63
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
    64
fun cabs_const (S, T) =
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
    65
  Const(@{const_name Abs_CFun}, (S --> T) --> (S ->> T));
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
    66
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
    67
fun mk_cabs t =
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
    68
  let val T = Term.fastype_of t
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
    69
  in cabs_const (Term.domain_type T, Term.range_type T) $ t end
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
    70
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
    71
(* builds the expression (LAM v. rhs) *)
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
    72
fun big_lambda v rhs =
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
    73
  cabs_const (Term.fastype_of v, Term.fastype_of rhs) $ Term.lambda v rhs;
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
    74
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
    75
(* builds the expression (LAM v1 v2 .. vn. rhs) *)
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
    76
fun big_lambdas [] rhs = rhs
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
    77
  | big_lambdas (v::vs) rhs = big_lambda v (big_lambdas vs rhs);
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
    78
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
    79
fun mk_capply (t, u) =
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
    80
  let val (S, T) =
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
    81
    case Term.fastype_of t of
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
    82
        Type(@{type_name "->"}, [S, T]) => (S, T)
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
    83
      | _ => 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
    84
  in capply_const (S, T) $ t $ u end;
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
    85
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
    86
(* miscellaneous term constructions *)
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
    87
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
    88
val mk_trp = HOLogic.mk_Trueprop;
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
    89
33775
7a1518c42c56 cleaned up; factored out fixed-point definition code
huffman
parents: 33774
diff changeset
    90
val mk_fst = HOLogic.mk_fst;
7a1518c42c56 cleaned up; factored out fixed-point definition code
huffman
parents: 33774
diff changeset
    91
val mk_snd = HOLogic.mk_snd;
7a1518c42c56 cleaned up; factored out fixed-point definition code
huffman
parents: 33774
diff changeset
    92
33774
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
    93
fun mk_cont t =
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
    94
  let val T = Term.fastype_of t
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
    95
  in Const(@{const_name cont}, T --> HOLogic.boolT) $ t end;
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
    96
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
    97
fun mk_fix t =
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
    98
  let val (T, _) = dest_cfunT (Term.fastype_of t)
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
    99
  in mk_capply (Const(@{const_name fix}, (T ->> T) ->> T), t) end;
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   100
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   101
fun mk_Rep_of T =
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   102
  Const (@{const_name Rep_of}, Term.itselfT T --> deflT) $ Logic.mk_type T;
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   103
33775
7a1518c42c56 cleaned up; factored out fixed-point definition code
huffman
parents: 33774
diff changeset
   104
(* 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
   105
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
   106
7a1518c42c56 cleaned up; factored out fixed-point definition code
huffman
parents: 33774
diff changeset
   107
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
   108
7a1518c42c56 cleaned up; factored out fixed-point definition code
huffman
parents: 33774
diff changeset
   109
(******************************************************************************)
7a1518c42c56 cleaned up; factored out fixed-point definition code
huffman
parents: 33774
diff changeset
   110
(*************** fixed-point definitions and unfolding theorems ***************)
7a1518c42c56 cleaned up; factored out fixed-point definition code
huffman
parents: 33774
diff changeset
   111
(******************************************************************************)
7a1518c42c56 cleaned up; factored out fixed-point definition code
huffman
parents: 33774
diff changeset
   112
7a1518c42c56 cleaned up; factored out fixed-point definition code
huffman
parents: 33774
diff changeset
   113
fun add_fixdefs
7a1518c42c56 cleaned up; factored out fixed-point definition code
huffman
parents: 33774
diff changeset
   114
    (spec : (binding * term) list)
7a1518c42c56 cleaned up; factored out fixed-point definition code
huffman
parents: 33774
diff changeset
   115
    (thy : theory) : thm list * theory =
7a1518c42c56 cleaned up; factored out fixed-point definition code
huffman
parents: 33774
diff changeset
   116
  let
7a1518c42c56 cleaned up; factored out fixed-point definition code
huffman
parents: 33774
diff changeset
   117
    val binds = map fst spec;
7a1518c42c56 cleaned up; factored out fixed-point definition code
huffman
parents: 33774
diff changeset
   118
    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
   119
    val functional = lambda_tuple lhss (mk_tuple rhss);
7a1518c42c56 cleaned up; factored out fixed-point definition code
huffman
parents: 33774
diff changeset
   120
    val fixpoint = mk_fix (mk_cabs functional);
7a1518c42c56 cleaned up; factored out fixed-point definition code
huffman
parents: 33774
diff changeset
   121
7a1518c42c56 cleaned up; factored out fixed-point definition code
huffman
parents: 33774
diff changeset
   122
    (* project components of fixpoint *)
7a1518c42c56 cleaned up; factored out fixed-point definition code
huffman
parents: 33774
diff changeset
   123
    fun mk_projs (x::[]) t = [(x, t)]
7a1518c42c56 cleaned up; factored out fixed-point definition code
huffman
parents: 33774
diff changeset
   124
      | 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
   125
    val projs = mk_projs lhss fixpoint;
7a1518c42c56 cleaned up; factored out fixed-point definition code
huffman
parents: 33774
diff changeset
   126
7a1518c42c56 cleaned up; factored out fixed-point definition code
huffman
parents: 33774
diff changeset
   127
    (* convert parameters to lambda abstractions *)
7a1518c42c56 cleaned up; factored out fixed-point definition code
huffman
parents: 33774
diff changeset
   128
    fun mk_eqn (lhs, rhs) =
7a1518c42c56 cleaned up; factored out fixed-point definition code
huffman
parents: 33774
diff changeset
   129
        case lhs of
7a1518c42c56 cleaned up; factored out fixed-point definition code
huffman
parents: 33774
diff changeset
   130
          Const (@{const_name Rep_CFun}, _) $ f $ (x as Free _) =>
7a1518c42c56 cleaned up; factored out fixed-point definition code
huffman
parents: 33774
diff changeset
   131
            mk_eqn (f, big_lambda x rhs)
7a1518c42c56 cleaned up; factored out fixed-point definition code
huffman
parents: 33774
diff changeset
   132
        | Const _ => Logic.mk_equals (lhs, rhs)
7a1518c42c56 cleaned up; factored out fixed-point definition code
huffman
parents: 33774
diff changeset
   133
        | _ => raise TERM ("lhs not of correct form", [lhs, rhs]);
7a1518c42c56 cleaned up; factored out fixed-point definition code
huffman
parents: 33774
diff changeset
   134
    val eqns = map mk_eqn projs;
7a1518c42c56 cleaned up; factored out fixed-point definition code
huffman
parents: 33774
diff changeset
   135
7a1518c42c56 cleaned up; factored out fixed-point definition code
huffman
parents: 33774
diff changeset
   136
    (* register constant definitions *)
33777
69eae9bca167 get rid of numbers on thy variables
huffman
parents: 33776
diff changeset
   137
    val (fixdef_thms, thy) =
33775
7a1518c42c56 cleaned up; factored out fixed-point definition code
huffman
parents: 33774
diff changeset
   138
      (PureThy.add_defs false o map Thm.no_attributes)
7a1518c42c56 cleaned up; factored out fixed-point definition code
huffman
parents: 33774
diff changeset
   139
        (map (Binding.suffix_name "_def") binds ~~ eqns) thy;
7a1518c42c56 cleaned up; factored out fixed-point definition code
huffman
parents: 33774
diff changeset
   140
7a1518c42c56 cleaned up; factored out fixed-point definition code
huffman
parents: 33774
diff changeset
   141
    (* prove applied version of definitions *)
7a1518c42c56 cleaned up; factored out fixed-point definition code
huffman
parents: 33774
diff changeset
   142
    fun prove_proj (lhs, rhs) =
7a1518c42c56 cleaned up; factored out fixed-point definition code
huffman
parents: 33774
diff changeset
   143
      let
7a1518c42c56 cleaned up; factored out fixed-point definition code
huffman
parents: 33774
diff changeset
   144
        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
   145
        val goal = Logic.mk_equals (lhs, rhs);
33777
69eae9bca167 get rid of numbers on thy variables
huffman
parents: 33776
diff changeset
   146
      in Goal.prove_global thy [] [] goal (K tac) end;
33775
7a1518c42c56 cleaned up; factored out fixed-point definition code
huffman
parents: 33774
diff changeset
   147
    val proj_thms = map prove_proj projs;
7a1518c42c56 cleaned up; factored out fixed-point definition code
huffman
parents: 33774
diff changeset
   148
7a1518c42c56 cleaned up; factored out fixed-point definition code
huffman
parents: 33774
diff changeset
   149
    (* mk_tuple lhss == fixpoint *)
7a1518c42c56 cleaned up; factored out fixed-point definition code
huffman
parents: 33774
diff changeset
   150
    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
   151
    val tuple_fixdef_thm = foldr1 pair_equalI proj_thms;
7a1518c42c56 cleaned up; factored out fixed-point definition code
huffman
parents: 33774
diff changeset
   152
7a1518c42c56 cleaned up; factored out fixed-point definition code
huffman
parents: 33774
diff changeset
   153
    val cont_thm =
33777
69eae9bca167 get rid of numbers on thy variables
huffman
parents: 33776
diff changeset
   154
      Goal.prove_global thy [] [] (mk_trp (mk_cont functional))
33775
7a1518c42c56 cleaned up; factored out fixed-point definition code
huffman
parents: 33774
diff changeset
   155
        (K (beta_tac 1));
7a1518c42c56 cleaned up; factored out fixed-point definition code
huffman
parents: 33774
diff changeset
   156
    val tuple_unfold_thm =
7a1518c42c56 cleaned up; factored out fixed-point definition code
huffman
parents: 33774
diff changeset
   157
      (@{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
   158
      |> LocalDefs.unfold (ProofContext.init thy) @{thms split_conv};
33775
7a1518c42c56 cleaned up; factored out fixed-point definition code
huffman
parents: 33774
diff changeset
   159
7a1518c42c56 cleaned up; factored out fixed-point definition code
huffman
parents: 33774
diff changeset
   160
    fun mk_unfold_thms [] thm = []
7a1518c42c56 cleaned up; factored out fixed-point definition code
huffman
parents: 33774
diff changeset
   161
      | mk_unfold_thms (n::[]) thm = [(n, thm)]
7a1518c42c56 cleaned up; factored out fixed-point definition code
huffman
parents: 33774
diff changeset
   162
      | mk_unfold_thms (n::ns) thm = let
7a1518c42c56 cleaned up; factored out fixed-point definition code
huffman
parents: 33774
diff changeset
   163
          val thmL = thm RS @{thm Pair_eqD1};
7a1518c42c56 cleaned up; factored out fixed-point definition code
huffman
parents: 33774
diff changeset
   164
          val thmR = thm RS @{thm Pair_eqD2};
7a1518c42c56 cleaned up; factored out fixed-point definition code
huffman
parents: 33774
diff changeset
   165
        in (n, thmL) :: mk_unfold_thms ns thmR end;
7a1518c42c56 cleaned up; factored out fixed-point definition code
huffman
parents: 33774
diff changeset
   166
    val unfold_binds = map (Binding.suffix_name "_unfold") binds;
7a1518c42c56 cleaned up; factored out fixed-point definition code
huffman
parents: 33774
diff changeset
   167
7a1518c42c56 cleaned up; factored out fixed-point definition code
huffman
parents: 33774
diff changeset
   168
    (* register unfold theorems *)
33777
69eae9bca167 get rid of numbers on thy variables
huffman
parents: 33776
diff changeset
   169
    val (unfold_thms, thy) =
33775
7a1518c42c56 cleaned up; factored out fixed-point definition code
huffman
parents: 33774
diff changeset
   170
      (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
   171
        (mk_unfold_thms unfold_binds tuple_unfold_thm) thy;
33775
7a1518c42c56 cleaned up; factored out fixed-point definition code
huffman
parents: 33774
diff changeset
   172
  in
33777
69eae9bca167 get rid of numbers on thy variables
huffman
parents: 33776
diff changeset
   173
    (unfold_thms, thy)
33775
7a1518c42c56 cleaned up; factored out fixed-point definition code
huffman
parents: 33774
diff changeset
   174
  end;
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
33774
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   177
(******************************************************************************)
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   178
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   179
fun typ_of_dtyp
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   180
    (descr : (string * string list) list)
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   181
    (sorts : (string * sort) list)
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   182
    : DatatypeAux.dtyp -> typ =
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   183
  let
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   184
    fun tfree a = TFree (a, the (AList.lookup (op =) sorts a))
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   185
    fun typ_of (DatatypeAux.DtTFree a) = tfree a
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   186
      | typ_of (DatatypeAux.DtType (s, ds)) = Type (s, map typ_of ds)
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   187
      | typ_of (DatatypeAux.DtRec i) =
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   188
          let val (s, vs) = nth descr i
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   189
          in Type (s, map tfree vs) end
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   190
  in typ_of end;
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   191
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   192
fun is_closed_dtyp (DatatypeAux.DtTFree a) = false
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   193
  | is_closed_dtyp (DatatypeAux.DtRec i) = false
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   194
  | is_closed_dtyp (DatatypeAux.DtType (s, ds)) = forall is_closed_dtyp ds;
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   195
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   196
(* FIXME: use theory data for this *)
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   197
val defl_tab : term Symtab.table =
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   198
    Symtab.make [(@{type_name "->"}, @{term "cfun_typ"}),
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   199
                 (@{type_name "++"}, @{term "ssum_typ"}),
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   200
                 (@{type_name "**"}, @{term "sprod_typ"}),
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   201
                 (@{type_name "*"}, @{term "cprod_typ"}),
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   202
                 (@{type_name "u"}, @{term "u_typ"}),
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   203
                 (@{type_name "upper_pd"}, @{term "upper_typ"}),
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   204
                 (@{type_name "lower_pd"}, @{term "lower_typ"}),
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   205
                 (@{type_name "convex_pd"}, @{term "convex_typ"})];
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   206
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   207
fun defl_of_typ
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   208
    (tab : term Symtab.table)
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   209
    (free : string -> term)
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   210
    (T : typ) : term =
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   211
  let
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   212
    fun is_closed_typ (Type (_, Ts)) = forall is_closed_typ Ts
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   213
      | is_closed_typ _ = false;
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   214
    fun defl_of (TFree (a, _)) = free a
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   215
      | defl_of (TVar _) = error ("defl_of_typ: TVar")
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   216
      | defl_of (T as Type (c, Ts)) =
33775
7a1518c42c56 cleaned up; factored out fixed-point definition code
huffman
parents: 33774
diff changeset
   217
        case Symtab.lookup tab c of
33774
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   218
          SOME t => Library.foldl mk_capply (t, map defl_of Ts)
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   219
        | NONE => if is_closed_typ T
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   220
                  then mk_Rep_of T
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   221
                  else error ("defl_of_typ: type variable under unsupported type constructor " ^ c);
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   222
  in defl_of T end;
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   223
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   224
fun defl_of_dtyp
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   225
    (descr : (string * string list) list)
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   226
    (sorts : (string * sort) list)
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   227
    (f : string -> term)
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   228
    (r : int -> term)
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   229
    (dt : DatatypeAux.dtyp) : term =
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   230
  let
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   231
    fun tfree a = TFree (a, the (AList.lookup (op =) sorts a))
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   232
    fun defl_of (DatatypeAux.DtTFree a) = f a
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   233
      | defl_of (DatatypeAux.DtRec i) = r i
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   234
      | defl_of (dt as DatatypeAux.DtType (s, ds)) =
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   235
        case Symtab.lookup defl_tab s of
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   236
          SOME t => Library.foldl mk_capply (t, map defl_of ds)
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   237
        | NONE => if DatatypeAux.is_rec_type dt
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   238
                  then error ("defl_of_dtyp: recursion under unsupported type constructor " ^ s)
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   239
                  else if is_closed_dtyp dt
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   240
                  then mk_Rep_of (typ_of_dtyp descr sorts dt)
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   241
                  else error ("defl_of_dtyp: type variable under unsupported type constructor " ^ s);
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   242
  in defl_of dt end;
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   243
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   244
(******************************************************************************)
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   245
(* prepare datatype specifications *)
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   246
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   247
fun read_typ thy str sorts =
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   248
  let
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   249
    val ctxt = ProofContext.init thy
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   250
      |> fold (Variable.declare_typ o TFree) sorts;
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   251
    val T = Syntax.read_typ ctxt str;
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   252
  in (T, Term.add_tfreesT T sorts) end;
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   253
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   254
fun cert_typ sign raw_T sorts =
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   255
  let
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   256
    val T = Type.no_tvars (Sign.certify_typ sign raw_T)
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   257
      handle TYPE (msg, _, _) => error msg;
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   258
    val sorts' = Term.add_tfreesT T sorts;
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   259
    val _ =
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   260
      case duplicates (op =) (map fst sorts') of
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   261
        [] => ()
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   262
      | dups => error ("Inconsistent sort constraints for " ^ commas dups)
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   263
  in (T, sorts') end;
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   264
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   265
fun gen_domain_isomorphism
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   266
    (prep_typ: theory -> 'a -> (string * sort) list -> typ * (string * sort) list)
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   267
    (doms_raw: (string list * binding * mixfix * 'a) list)
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   268
    (thy: theory)
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   269
    : theory =
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   270
  let
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   271
    val _ = Theory.requires thy "Domain" "domain definitions";
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   272
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   273
    (* this theory is used just for parsing *)
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   274
    val tmp_thy = thy |>
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   275
      Theory.copy |>
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   276
      Sign.add_types (map (fn (tvs, tname, mx, _) =>
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   277
        (tname, length tvs, mx)) doms_raw);
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   278
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   279
    fun prep_dom thy (vs, t, mx, typ_raw) sorts =
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   280
      let val (typ, sorts') = prep_typ thy typ_raw sorts
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   281
      in ((vs, t, mx, typ), sorts') end;
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   282
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   283
    val (doms : (string list * binding * mixfix * typ) list,
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   284
         sorts : (string * sort) list) =
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   285
      fold_map (prep_dom tmp_thy) doms_raw [];
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   286
33775
7a1518c42c56 cleaned up; factored out fixed-point definition code
huffman
parents: 33774
diff changeset
   287
    (* domain equations *)
7a1518c42c56 cleaned up; factored out fixed-point definition code
huffman
parents: 33774
diff changeset
   288
    fun mk_dom_eqn (vs, tbind, mx, rhs) =
7a1518c42c56 cleaned up; factored out fixed-point definition code
huffman
parents: 33774
diff changeset
   289
      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
   290
      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
   291
    val dom_eqns = map mk_dom_eqn doms;
7a1518c42c56 cleaned up; factored out fixed-point definition code
huffman
parents: 33774
diff changeset
   292
7a1518c42c56 cleaned up; factored out fixed-point definition code
huffman
parents: 33774
diff changeset
   293
    (* check for valid type parameters *)
33774
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   294
    val (tyvars, _, _, _)::_ = doms;
33775
7a1518c42c56 cleaned up; factored out fixed-point definition code
huffman
parents: 33774
diff changeset
   295
    val new_doms = map (fn (tvs, tname, mx, _) =>
33774
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   296
      let val full_tname = Sign.full_name tmp_thy tname
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   297
      in
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   298
        (case duplicates (op =) tvs of
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   299
          [] =>
33775
7a1518c42c56 cleaned up; factored out fixed-point definition code
huffman
parents: 33774
diff changeset
   300
            if eq_set (op =) (tyvars, tvs) then (full_tname, tvs)
33774
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   301
            else error ("Mutually recursive domains must have same type parameters")
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   302
        | dups => error ("Duplicate parameter(s) for domain " ^ quote (Binding.str_of tname) ^
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   303
            " : " ^ commas dups))
33775
7a1518c42c56 cleaned up; factored out fixed-point definition code
huffman
parents: 33774
diff changeset
   304
      end) doms;
33774
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   305
    val dom_names = map fst new_doms;
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   306
33775
7a1518c42c56 cleaned up; factored out fixed-point definition code
huffman
parents: 33774
diff changeset
   307
    (* declare type combinator constants *)
7a1518c42c56 cleaned up; factored out fixed-point definition code
huffman
parents: 33774
diff changeset
   308
    fun declare_typ_const (vs, tbind, mx, rhs) thy =
33774
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   309
      let
33775
7a1518c42c56 cleaned up; factored out fixed-point definition code
huffman
parents: 33774
diff changeset
   310
        val typ_type = Library.foldr cfunT (map (K deflT) vs, deflT);
33774
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   311
        val typ_bind = Binding.suffix_name "_typ" tbind;
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   312
      in
33775
7a1518c42c56 cleaned up; factored out fixed-point definition code
huffman
parents: 33774
diff changeset
   313
        Sign.declare_const ((typ_bind, typ_type), NoSyn) thy
33774
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   314
      end;
33777
69eae9bca167 get rid of numbers on thy variables
huffman
parents: 33776
diff changeset
   315
    val (typ_consts, thy) = fold_map declare_typ_const doms thy;
33774
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   316
33775
7a1518c42c56 cleaned up; factored out fixed-point definition code
huffman
parents: 33774
diff changeset
   317
    (* defining equations for type combinators *)
7a1518c42c56 cleaned up; factored out fixed-point definition code
huffman
parents: 33774
diff changeset
   318
    val defl_tab1 = defl_tab; (* FIXME: use theory data *)
7a1518c42c56 cleaned up; factored out fixed-point definition code
huffman
parents: 33774
diff changeset
   319
    val defl_tab2 =
7a1518c42c56 cleaned up; factored out fixed-point definition code
huffman
parents: 33774
diff changeset
   320
      Symtab.make (map (fst o dest_Type o fst) dom_eqns ~~ typ_consts);
7a1518c42c56 cleaned up; factored out fixed-point definition code
huffman
parents: 33774
diff changeset
   321
    val defl_tab' = Symtab.merge (K true) (defl_tab1, defl_tab2);
7a1518c42c56 cleaned up; factored out fixed-point definition code
huffman
parents: 33774
diff changeset
   322
    fun free a = Free (Library.unprefix "'" a, deflT);
33776
5048b02c2bbb automate proofs of REP equations
huffman
parents: 33775
diff changeset
   323
    fun mk_defl_spec (lhsT, rhsT) =
5048b02c2bbb automate proofs of REP equations
huffman
parents: 33775
diff changeset
   324
      mk_eqs (defl_of_typ defl_tab' free lhsT,
5048b02c2bbb automate proofs of REP equations
huffman
parents: 33775
diff changeset
   325
              defl_of_typ defl_tab' free rhsT);
33775
7a1518c42c56 cleaned up; factored out fixed-point definition code
huffman
parents: 33774
diff changeset
   326
    val defl_specs = map mk_defl_spec dom_eqns;
33774
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   327
33775
7a1518c42c56 cleaned up; factored out fixed-point definition code
huffman
parents: 33774
diff changeset
   328
    (* register recursive definition of type combinators *)
7a1518c42c56 cleaned up; factored out fixed-point definition code
huffman
parents: 33774
diff changeset
   329
    val typ_binds = map (Binding.suffix_name "_typ" o #2) doms;
33777
69eae9bca167 get rid of numbers on thy variables
huffman
parents: 33776
diff changeset
   330
    val (typ_unfold_thms, thy) = add_fixdefs (typ_binds ~~ defl_specs) thy;
33774
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   331
33776
5048b02c2bbb automate proofs of REP equations
huffman
parents: 33775
diff changeset
   332
    (* define types using deflation combinators *)
33774
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   333
    fun make_repdef ((vs, tbind, mx, _), typ_const) thy =
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   334
      let
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   335
        fun tfree a = TFree (a, the (AList.lookup (op =) sorts a))
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   336
        val reps = map (mk_Rep_of o tfree) vs;
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   337
        val defl = Library.foldl mk_capply (typ_const, reps);
33777
69eae9bca167 get rid of numbers on thy variables
huffman
parents: 33776
diff changeset
   338
        val ((_, _, _, {REP, ...}), thy) =
33774
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   339
          Repdef.add_repdef false NONE (tbind, vs, mx) defl NONE thy;
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   340
      in
33777
69eae9bca167 get rid of numbers on thy variables
huffman
parents: 33776
diff changeset
   341
        (REP, thy)
33774
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   342
      end;
33777
69eae9bca167 get rid of numbers on thy variables
huffman
parents: 33776
diff changeset
   343
    val (REP_thms, thy) =
69eae9bca167 get rid of numbers on thy variables
huffman
parents: 33776
diff changeset
   344
      fold_map make_repdef (doms ~~ typ_consts) thy;
33774
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   345
33776
5048b02c2bbb automate proofs of REP equations
huffman
parents: 33775
diff changeset
   346
    (* FIXME: use theory data for this *)
5048b02c2bbb automate proofs of REP equations
huffman
parents: 33775
diff changeset
   347
    val REP_simps = REP_thms @
5048b02c2bbb automate proofs of REP equations
huffman
parents: 33775
diff changeset
   348
      @{thms REP_cfun REP_ssum REP_sprod REP_cprod REP_up
5048b02c2bbb automate proofs of REP equations
huffman
parents: 33775
diff changeset
   349
             REP_upper REP_lower REP_convex};
5048b02c2bbb automate proofs of REP equations
huffman
parents: 33775
diff changeset
   350
5048b02c2bbb automate proofs of REP equations
huffman
parents: 33775
diff changeset
   351
    (* prove REP equations *)
5048b02c2bbb automate proofs of REP equations
huffman
parents: 33775
diff changeset
   352
    fun mk_REP_eqn_thm (lhsT, rhsT) =
5048b02c2bbb automate proofs of REP equations
huffman
parents: 33775
diff changeset
   353
      let
5048b02c2bbb automate proofs of REP equations
huffman
parents: 33775
diff changeset
   354
        val goal = mk_eqs (mk_Rep_of lhsT, mk_Rep_of rhsT);
5048b02c2bbb automate proofs of REP equations
huffman
parents: 33775
diff changeset
   355
        val tac =
5048b02c2bbb automate proofs of REP equations
huffman
parents: 33775
diff changeset
   356
          simp_tac (HOL_basic_ss addsimps REP_simps) 1
5048b02c2bbb automate proofs of REP equations
huffman
parents: 33775
diff changeset
   357
          THEN resolve_tac typ_unfold_thms 1;
5048b02c2bbb automate proofs of REP equations
huffman
parents: 33775
diff changeset
   358
      in
33777
69eae9bca167 get rid of numbers on thy variables
huffman
parents: 33776
diff changeset
   359
        Goal.prove_global thy [] [] goal (K tac)
33776
5048b02c2bbb automate proofs of REP equations
huffman
parents: 33775
diff changeset
   360
      end;
5048b02c2bbb automate proofs of REP equations
huffman
parents: 33775
diff changeset
   361
    val REP_eqn_thms = map mk_REP_eqn_thm dom_eqns;
5048b02c2bbb automate proofs of REP equations
huffman
parents: 33775
diff changeset
   362
33774
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   363
  in
33777
69eae9bca167 get rid of numbers on thy variables
huffman
parents: 33776
diff changeset
   364
    thy
33774
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   365
  end;
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   366
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   367
val domain_isomorphism = gen_domain_isomorphism cert_typ;
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   368
val domain_isomorphism_cmd = gen_domain_isomorphism read_typ;
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   369
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   370
(******************************************************************************)
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   371
(******************************** outer syntax ********************************)
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   372
(******************************************************************************)
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   373
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   374
local
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   375
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   376
structure P = OuterParse and K = OuterKeyword
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   377
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   378
val parse_domain_iso : (string list * binding * mixfix * string) parser =
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   379
  (P.type_args -- P.binding -- P.opt_infix -- (P.$$$ "=" |-- P.typ))
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   380
    >> (fn (((vs, t), mx), rhs) => (vs, t, mx, rhs));
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   381
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   382
val parse_domain_isos = P.and_list1 parse_domain_iso;
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   383
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   384
in
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   385
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   386
val _ =
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   387
  OuterSyntax.command "domain_isomorphism" "define domain isomorphisms (HOLCF)" K.thy_decl
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   388
    (parse_domain_isos >> (Toplevel.theory o domain_isomorphism_cmd));
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   389
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   390
end;
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   391
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   392
end;