Nominal/nominal_function_core.ML
author Christian Urban <urbanc@in.tum.de>
Thu, 19 Apr 2018 13:57:17 +0100
changeset 3245 017e33849f4d
parent 3244 a44479bde681
permissions -rw-r--r--
updated to Isabelle 2016-1
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
2665
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     1
(*  Nominal Function Core
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     2
    Author:  Christian Urban
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     3
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     4
    heavily based on the code of Alexander Krauss
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     5
    (code forked on 14 January 2011)
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     6
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     7
Core of the nominal function package.
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     8
*)
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     9
2821
c7d4bd9e89e0 fixed problem with earlier commit about nominal_function_common; added facility for specifying an invariant - added a definition of frees_set which need a finiteness invariant
Christian Urban <urbanc@in.tum.de>
parents: 2819
diff changeset
    10
2665
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    11
signature NOMINAL_FUNCTION_CORE =
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    12
sig
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    13
  val trace: bool Unsynchronized.ref
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    14
2819
4bd584ff4fab added an option for an invariant (at the moment only a stub)
Christian Urban <urbanc@in.tum.de>
parents: 2803
diff changeset
    15
  val prepare_nominal_function : Nominal_Function_Common.nominal_function_config
2665
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    16
    -> string (* defname *)
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    17
    -> ((bstring * typ) * mixfix) list (* defined symbol *)
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    18
    -> ((bstring * typ) list * term list * term * term) list (* specification *)
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    19
    -> local_theory
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    20
    -> (term   (* f *)
3197
25d11b449e92 definition of an auxiliary graph in nominal-primrec definitions
Christian Urban <urbanc@in.tum.de>
parents: 3108
diff changeset
    21
        * term (* G(raph) *)
25d11b449e92 definition of an auxiliary graph in nominal-primrec definitions
Christian Urban <urbanc@in.tum.de>
parents: 3108
diff changeset
    22
        * thm list (* GIntros *)
25d11b449e92 definition of an auxiliary graph in nominal-primrec definitions
Christian Urban <urbanc@in.tum.de>
parents: 3108
diff changeset
    23
        * thm (* Ginduct *) 
2665
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    24
        * thm  (* goalstate *)
3244
a44479bde681 fixed a problem with two example theories
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3239
diff changeset
    25
        * (Proof.context -> thm -> Nominal_Function_Common.nominal_function_result) (* continuation *)
2665
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    26
       ) * local_theory
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    27
3197
25d11b449e92 definition of an auxiliary graph in nominal-primrec definitions
Christian Urban <urbanc@in.tum.de>
parents: 3108
diff changeset
    28
  val inductive_def : (binding * typ) * mixfix -> term list -> local_theory 
25d11b449e92 definition of an auxiliary graph in nominal-primrec definitions
Christian Urban <urbanc@in.tum.de>
parents: 3108
diff changeset
    29
    -> (term * thm list * thm * thm) * local_theory
2665
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    30
end
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    31
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    32
structure Nominal_Function_Core : NOMINAL_FUNCTION_CORE =
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    33
struct
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    34
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    35
val trace = Unsynchronized.ref false
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    36
fun trace_msg msg = if ! trace then tracing (msg ()) else ()
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    37
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    38
val boolT = HOLogic.boolT
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    39
val mk_eq = HOLogic.mk_eq
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    40
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    41
open Function_Lib
2821
c7d4bd9e89e0 fixed problem with earlier commit about nominal_function_common; added facility for specifying an invariant - added a definition of frees_set which need a finiteness invariant
Christian Urban <urbanc@in.tum.de>
parents: 2819
diff changeset
    42
open Function_Common
2819
4bd584ff4fab added an option for an invariant (at the moment only a stub)
Christian Urban <urbanc@in.tum.de>
parents: 2803
diff changeset
    43
open Nominal_Function_Common
2665
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    44
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    45
datatype globals = Globals of
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    46
 {fvar: term,
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    47
  domT: typ,
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    48
  ranT: typ,
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    49
  h: term,
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    50
  y: term,
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    51
  x: term,
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    52
  z: term,
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    53
  a: term,
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    54
  P: term,
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    55
  D: term,
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    56
  Pbool:term}
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    57
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    58
datatype rec_call_info = RCInfo of
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    59
 {RIvs: (string * typ) list,  (* Call context: fixes and assumes *)
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    60
  CCas: thm list,
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    61
  rcarg: term,                 (* The recursive argument *)
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    62
  llRI: thm,
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    63
  h_assum: term}
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    64
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    65
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    66
datatype clause_context = ClauseContext of
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    67
 {ctxt : Proof.context,
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    68
  qs : term list,
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    69
  gs : term list,
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    70
  lhs: term,
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    71
  rhs: term,
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    72
  cqs: cterm list,
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    73
  ags: thm list,
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    74
  case_hyp : thm}
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    75
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    76
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    77
fun transfer_clause_ctx thy (ClauseContext { ctxt, qs, gs, lhs, rhs, cqs, ags, case_hyp }) =
3045
d0ad264f8c4f updated to Isabelle 3 Nov; it includes a hack to work around a bug in the localised version of the quotient package
Christian Urban <urbanc@in.tum.de>
parents: 2994
diff changeset
    78
  ClauseContext { ctxt = Proof_Context.transfer thy ctxt,
2665
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    79
    qs = qs, gs = gs, lhs = lhs, rhs = rhs, cqs = cqs, ags = ags, case_hyp = case_hyp }
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    80
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    81
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    82
datatype clause_info = ClauseInfo of
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    83
 {no: int,
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    84
  qglr : ((string * typ) list * term list * term * term),
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    85
  cdata : clause_context,
3239
67370521c09c updated for Isabelle 2015
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3229
diff changeset
    86
  tree: Function_Context_Tree.ctx_tree,
2665
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    87
  lGI: thm,
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    88
  RCs: rec_call_info list}
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    89
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    90
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    91
(* Theory dependencies. *)
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    92
val acc_induct_rule = @{thm accp_induct_rule}
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    93
3229
b52e8651591f updated to Isabelle changes
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3227
diff changeset
    94
val ex1_implies_ex = @{thm Fun_Def.fundef_ex1_existence}
b52e8651591f updated to Isabelle changes
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3227
diff changeset
    95
val ex1_implies_un = @{thm Fun_Def.fundef_ex1_uniqueness}
b52e8651591f updated to Isabelle changes
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3227
diff changeset
    96
val ex1_implies_iff = @{thm Fun_Def.fundef_ex1_iff}
2665
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    97
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    98
val acc_downward = @{thm accp_downward}
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    99
val accI = @{thm accp.accI}
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   100
val case_split = @{thm HOL.case_split}
3229
b52e8651591f updated to Isabelle changes
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3227
diff changeset
   101
val fundef_default_value = @{thm Fun_Def.fundef_default_value}
2665
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   102
val not_acc_down = @{thm not_accp_down}
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   103
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   104
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   105
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   106
fun find_calls tree =
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   107
  let
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   108
    fun add_Ri (fixes,assumes) (_ $ arg) _ (_, xs) =
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   109
      ([], (fixes, assumes, arg) :: xs)
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   110
      | add_Ri _ _ _ _ = raise Match
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   111
  in
3239
67370521c09c updated for Isabelle 2015
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3229
diff changeset
   112
    rev (Function_Context_Tree.traverse_tree add_Ri tree [])
2665
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   113
  end
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   114
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   115
(* nominal *)
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   116
fun mk_eqvt_at (f_trm, arg_trm) =
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   117
  let
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   118
    val f_ty = fastype_of f_trm
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   119
    val arg_ty = domain_type f_ty
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   120
  in
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   121
    Const (@{const_name eqvt_at}, [f_ty, arg_ty] ---> @{typ bool}) $ f_trm $ arg_trm
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   122
    |> HOLogic.mk_Trueprop
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   123
  end
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   124
2707
747ebf2f066d made eqvt-proof explicit in the function definitions
Christian Urban <urbanc@in.tum.de>
parents: 2677
diff changeset
   125
fun mk_eqvt trm =
747ebf2f066d made eqvt-proof explicit in the function definitions
Christian Urban <urbanc@in.tum.de>
parents: 2677
diff changeset
   126
  let
747ebf2f066d made eqvt-proof explicit in the function definitions
Christian Urban <urbanc@in.tum.de>
parents: 2677
diff changeset
   127
    val ty = fastype_of trm
747ebf2f066d made eqvt-proof explicit in the function definitions
Christian Urban <urbanc@in.tum.de>
parents: 2677
diff changeset
   128
  in
747ebf2f066d made eqvt-proof explicit in the function definitions
Christian Urban <urbanc@in.tum.de>
parents: 2677
diff changeset
   129
    Const (@{const_name eqvt}, ty --> @{typ bool}) $ trm
747ebf2f066d made eqvt-proof explicit in the function definitions
Christian Urban <urbanc@in.tum.de>
parents: 2677
diff changeset
   130
    |> HOLogic.mk_Trueprop
747ebf2f066d made eqvt-proof explicit in the function definitions
Christian Urban <urbanc@in.tum.de>
parents: 2677
diff changeset
   131
  end
747ebf2f066d made eqvt-proof explicit in the function definitions
Christian Urban <urbanc@in.tum.de>
parents: 2677
diff changeset
   132
2821
c7d4bd9e89e0 fixed problem with earlier commit about nominal_function_common; added facility for specifying an invariant - added a definition of frees_set which need a finiteness invariant
Christian Urban <urbanc@in.tum.de>
parents: 2819
diff changeset
   133
fun mk_inv inv (f_trm, arg_trm) = 
c7d4bd9e89e0 fixed problem with earlier commit about nominal_function_common; added facility for specifying an invariant - added a definition of frees_set which need a finiteness invariant
Christian Urban <urbanc@in.tum.de>
parents: 2819
diff changeset
   134
  betapplys (inv, [arg_trm, (f_trm $ arg_trm)])
c7d4bd9e89e0 fixed problem with earlier commit about nominal_function_common; added facility for specifying an invariant - added a definition of frees_set which need a finiteness invariant
Christian Urban <urbanc@in.tum.de>
parents: 2819
diff changeset
   135
  |> HOLogic.mk_Trueprop
c7d4bd9e89e0 fixed problem with earlier commit about nominal_function_common; added facility for specifying an invariant - added a definition of frees_set which need a finiteness invariant
Christian Urban <urbanc@in.tum.de>
parents: 2819
diff changeset
   136
c7d4bd9e89e0 fixed problem with earlier commit about nominal_function_common; added facility for specifying an invariant - added a definition of frees_set which need a finiteness invariant
Christian Urban <urbanc@in.tum.de>
parents: 2819
diff changeset
   137
fun mk_invariant (Globals {x, y, ...}) G invariant = 
c7d4bd9e89e0 fixed problem with earlier commit about nominal_function_common; added facility for specifying an invariant - added a definition of frees_set which need a finiteness invariant
Christian Urban <urbanc@in.tum.de>
parents: 2819
diff changeset
   138
  let
c7d4bd9e89e0 fixed problem with earlier commit about nominal_function_common; added facility for specifying an invariant - added a definition of frees_set which need a finiteness invariant
Christian Urban <urbanc@in.tum.de>
parents: 2819
diff changeset
   139
    val prem = HOLogic.mk_Trueprop (G $ x $ y)
c7d4bd9e89e0 fixed problem with earlier commit about nominal_function_common; added facility for specifying an invariant - added a definition of frees_set which need a finiteness invariant
Christian Urban <urbanc@in.tum.de>
parents: 2819
diff changeset
   140
    val concl = HOLogic.mk_Trueprop (betapplys (invariant, [x, y]))
c7d4bd9e89e0 fixed problem with earlier commit about nominal_function_common; added facility for specifying an invariant - added a definition of frees_set which need a finiteness invariant
Christian Urban <urbanc@in.tum.de>
parents: 2819
diff changeset
   141
  in
c7d4bd9e89e0 fixed problem with earlier commit about nominal_function_common; added facility for specifying an invariant - added a definition of frees_set which need a finiteness invariant
Christian Urban <urbanc@in.tum.de>
parents: 2819
diff changeset
   142
    Logic.mk_implies (prem, concl)
c7d4bd9e89e0 fixed problem with earlier commit about nominal_function_common; added facility for specifying an invariant - added a definition of frees_set which need a finiteness invariant
Christian Urban <urbanc@in.tum.de>
parents: 2819
diff changeset
   143
    |> mk_forall_rename ("y", y)
c7d4bd9e89e0 fixed problem with earlier commit about nominal_function_common; added facility for specifying an invariant - added a definition of frees_set which need a finiteness invariant
Christian Urban <urbanc@in.tum.de>
parents: 2819
diff changeset
   144
    |> mk_forall_rename ("x", x)
c7d4bd9e89e0 fixed problem with earlier commit about nominal_function_common; added facility for specifying an invariant - added a definition of frees_set which need a finiteness invariant
Christian Urban <urbanc@in.tum.de>
parents: 2819
diff changeset
   145
  end  
c7d4bd9e89e0 fixed problem with earlier commit about nominal_function_common; added facility for specifying an invariant - added a definition of frees_set which need a finiteness invariant
Christian Urban <urbanc@in.tum.de>
parents: 2819
diff changeset
   146
2790
d2154b421707 tuned last commit
Christian Urban <urbanc@in.tum.de>
parents: 2789
diff changeset
   147
(** building proof obligations *)
2802
3b9ef98a03d2 first test to fix the problem with free variables
Christian Urban <urbanc@in.tum.de>
parents: 2796
diff changeset
   148
fun mk_eqvt_proof_obligation qs fvar (vs, assms, arg) = 
2789
32979078bfe9 functions involving if and case do not throw exceptions anymore; but eqvt_at assumption has now a precondition
Christian Urban <urbanc@in.tum.de>
parents: 2788
diff changeset
   149
  mk_eqvt_at (fvar, arg)
3239
67370521c09c updated for Isabelle 2015
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3229
diff changeset
   150
  |> curry Logic.list_implies (map Thm.prop_of assms)
3108
61db5ad429bb updated to Isabelle 16 January
Christian Urban <urbanc@in.tum.de>
parents: 3045
diff changeset
   151
  |> fold_rev (Logic.all o Free) vs
2994
4ee772b12032 Update to new Isabelle
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2974
diff changeset
   152
  |> fold_rev absfree qs
2789
32979078bfe9 functions involving if and case do not throw exceptions anymore; but eqvt_at assumption has now a precondition
Christian Urban <urbanc@in.tum.de>
parents: 2788
diff changeset
   153
  |> strip_abs_body
2665
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   154
2821
c7d4bd9e89e0 fixed problem with earlier commit about nominal_function_common; added facility for specifying an invariant - added a definition of frees_set which need a finiteness invariant
Christian Urban <urbanc@in.tum.de>
parents: 2819
diff changeset
   155
fun mk_inv_proof_obligation inv qs fvar (vs, assms, arg) = 
c7d4bd9e89e0 fixed problem with earlier commit about nominal_function_common; added facility for specifying an invariant - added a definition of frees_set which need a finiteness invariant
Christian Urban <urbanc@in.tum.de>
parents: 2819
diff changeset
   156
  mk_inv inv (fvar, arg)
3239
67370521c09c updated for Isabelle 2015
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3229
diff changeset
   157
  |> curry Logic.list_implies (map Thm.prop_of assms)
3108
61db5ad429bb updated to Isabelle 16 January
Christian Urban <urbanc@in.tum.de>
parents: 3045
diff changeset
   158
  |> fold_rev (Logic.all o Free) vs
2994
4ee772b12032 Update to new Isabelle
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2974
diff changeset
   159
  |> fold_rev absfree qs
2821
c7d4bd9e89e0 fixed problem with earlier commit about nominal_function_common; added facility for specifying an invariant - added a definition of frees_set which need a finiteness invariant
Christian Urban <urbanc@in.tum.de>
parents: 2819
diff changeset
   160
  |> strip_abs_body
c7d4bd9e89e0 fixed problem with earlier commit about nominal_function_common; added facility for specifying an invariant - added a definition of frees_set which need a finiteness invariant
Christian Urban <urbanc@in.tum.de>
parents: 2819
diff changeset
   161
2665
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   162
(** building proof obligations *)
2821
c7d4bd9e89e0 fixed problem with earlier commit about nominal_function_common; added facility for specifying an invariant - added a definition of frees_set which need a finiteness invariant
Christian Urban <urbanc@in.tum.de>
parents: 2819
diff changeset
   163
fun mk_compat_proof_obligations domT ranT fvar f RCss inv glrs =
2665
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   164
  let
2862
47063163f333 added eqvt_at and invariant for boths sides of the equations
Christian Urban <urbanc@in.tum.de>
parents: 2848
diff changeset
   165
    fun mk_impl (((qs, gs, lhs, rhs), RCs_lhs), ((qs', gs', lhs', rhs'), RCs_rhs)) =
2665
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   166
      let
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   167
        val shift = incr_boundvars (length qs')
2862
47063163f333 added eqvt_at and invariant for boths sides of the equations
Christian Urban <urbanc@in.tum.de>
parents: 2848
diff changeset
   168
        val eqvts_obligations_lhs = map (shift o mk_eqvt_proof_obligation qs fvar) RCs_lhs
47063163f333 added eqvt_at and invariant for boths sides of the equations
Christian Urban <urbanc@in.tum.de>
parents: 2848
diff changeset
   169
        val eqvts_obligations_rhs = map (mk_eqvt_proof_obligation qs' fvar) RCs_rhs
47063163f333 added eqvt_at and invariant for boths sides of the equations
Christian Urban <urbanc@in.tum.de>
parents: 2848
diff changeset
   170
        val invs_obligations_lhs = map (shift o mk_inv_proof_obligation inv qs fvar) RCs_lhs
47063163f333 added eqvt_at and invariant for boths sides of the equations
Christian Urban <urbanc@in.tum.de>
parents: 2848
diff changeset
   171
        val invs_obligations_rhs = map (mk_inv_proof_obligation inv qs' fvar) RCs_rhs
2665
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   172
      in
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   173
        Logic.mk_implies
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   174
          (HOLogic.mk_Trueprop (HOLogic.eq_const domT $ shift lhs $ lhs'),
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   175
            HOLogic.mk_Trueprop (HOLogic.eq_const ranT $ shift rhs $ rhs'))
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   176
        |> fold_rev (curry Logic.mk_implies) (map shift gs @ gs')
2862
47063163f333 added eqvt_at and invariant for boths sides of the equations
Christian Urban <urbanc@in.tum.de>
parents: 2848
diff changeset
   177
        |> fold_rev (curry Logic.mk_implies) invs_obligations_rhs (* nominal *)
47063163f333 added eqvt_at and invariant for boths sides of the equations
Christian Urban <urbanc@in.tum.de>
parents: 2848
diff changeset
   178
        |> fold_rev (curry Logic.mk_implies) invs_obligations_lhs (* nominal *)
47063163f333 added eqvt_at and invariant for boths sides of the equations
Christian Urban <urbanc@in.tum.de>
parents: 2848
diff changeset
   179
        |> fold_rev (curry Logic.mk_implies) eqvts_obligations_rhs (* nominal *)
47063163f333 added eqvt_at and invariant for boths sides of the equations
Christian Urban <urbanc@in.tum.de>
parents: 2848
diff changeset
   180
        |> fold_rev (curry Logic.mk_implies) eqvts_obligations_lhs (* nominal *)
3108
61db5ad429bb updated to Isabelle 16 January
Christian Urban <urbanc@in.tum.de>
parents: 3045
diff changeset
   181
        |> fold_rev (fn (n,T) => fn b => Logic.all_const T $ Abs(n,T,b)) (qs @ qs')
2665
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   182
        |> curry abstract_over fvar
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   183
        |> curry subst_bound f
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   184
      end
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   185
  in
2789
32979078bfe9 functions involving if and case do not throw exceptions anymore; but eqvt_at assumption has now a precondition
Christian Urban <urbanc@in.tum.de>
parents: 2788
diff changeset
   186
    map mk_impl (unordered_pairs (glrs ~~ RCss))
2665
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   187
  end
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   188
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   189
fun mk_completeness (Globals {x, Pbool, ...}) clauses qglrs =
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   190
  let
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   191
    fun mk_case (ClauseContext {qs, gs, lhs, ...}, (oqs, _, _, _)) =
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   192
      HOLogic.mk_Trueprop Pbool
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   193
      |> curry Logic.mk_implies (HOLogic.mk_Trueprop (mk_eq (x, lhs)))
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   194
      |> fold_rev (curry Logic.mk_implies) gs
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   195
      |> fold_rev mk_forall_rename (map fst oqs ~~ qs)
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   196
  in
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   197
    HOLogic.mk_Trueprop Pbool
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   198
    |> fold_rev (curry Logic.mk_implies o mk_case) (clauses ~~ qglrs)
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   199
    |> mk_forall_rename ("x", x)
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   200
    |> mk_forall_rename ("P", Pbool)
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   201
  end
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   202
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   203
(** making a context with it's own local bindings **)
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   204
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   205
fun mk_clause_context x ctxt (pre_qs,pre_gs,pre_lhs,pre_rhs) =
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   206
  let
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   207
    val (qs, ctxt') = Variable.variant_fixes (map fst pre_qs) ctxt
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   208
      |>> map2 (fn (_, T) => fn n => Free (n, T)) pre_qs
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   209
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   210
    fun inst t = subst_bounds (rev qs, t)
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   211
    val gs = map inst pre_gs
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   212
    val lhs = inst pre_lhs
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   213
    val rhs = inst pre_rhs
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   214
3239
67370521c09c updated for Isabelle 2015
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3229
diff changeset
   215
    val cqs = map (Thm.cterm_of ctxt') qs
67370521c09c updated for Isabelle 2015
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3229
diff changeset
   216
    val ags = map (Thm.assume o Thm.cterm_of ctxt') gs
2665
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   217
3239
67370521c09c updated for Isabelle 2015
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3229
diff changeset
   218
    val case_hyp = Thm.assume (Thm.cterm_of ctxt' (HOLogic.mk_Trueprop (mk_eq (x, lhs))))
2665
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   219
  in
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   220
    ClauseContext { ctxt = ctxt', qs = qs, gs = gs, lhs = lhs, rhs = rhs,
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   221
      cqs = cqs, ags = ags, case_hyp = case_hyp }
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   222
  end
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   223
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   224
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   225
(* lowlevel term function. FIXME: remove *)
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   226
fun abstract_over_list vs body =
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   227
  let
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   228
    fun abs lev v tm =
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   229
      if v aconv tm then Bound lev
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   230
      else
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   231
        (case tm of
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   232
          Abs (a, T, t) => Abs (a, T, abs (lev + 1) v t)
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   233
        | t $ u => abs lev v t $ abs lev v u
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   234
        | t => t)
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   235
  in
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   236
    fold_index (fn (i, v) => fn t => abs i v t) vs body
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   237
  end
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   238
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   239
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   240
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   241
fun mk_clause_info globals G f no cdata qglr tree RCs GIntro_thm RIntro_thms =
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   242
  let
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   243
    val Globals {h, ...} = globals
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   244
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   245
    val ClauseContext { ctxt, qs, cqs, ags, ... } = cdata
3239
67370521c09c updated for Isabelle 2015
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3229
diff changeset
   246
    val cert = Thm.cterm_of ctxt
2665
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   247
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   248
    (* Instantiate the GIntro thm with "f" and import into the clause context. *)
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   249
    val lGI = GIntro_thm
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   250
      |> Thm.forall_elim (cert f)
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   251
      |> fold Thm.forall_elim cqs
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   252
      |> fold Thm.elim_implies ags
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   253
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   254
    fun mk_call_info (rcfix, rcassm, rcarg) RI =
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   255
      let
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   256
        val llRI = RI
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   257
          |> fold Thm.forall_elim cqs
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   258
          |> fold (Thm.forall_elim o cert o Free) rcfix
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   259
          |> fold Thm.elim_implies ags
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   260
          |> fold Thm.elim_implies rcassm
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   261
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   262
        val h_assum =
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   263
          HOLogic.mk_Trueprop (G $ rcarg $ (h $ rcarg))
3239
67370521c09c updated for Isabelle 2015
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3229
diff changeset
   264
          |> fold_rev (curry Logic.mk_implies o Thm.prop_of) rcassm
2665
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   265
          |> fold_rev (Logic.all o Free) rcfix
3045
d0ad264f8c4f updated to Isabelle 3 Nov; it includes a hack to work around a bug in the localised version of the quotient package
Christian Urban <urbanc@in.tum.de>
parents: 2994
diff changeset
   266
          |> Pattern.rewrite_term (Proof_Context.theory_of ctxt) [(f, h)] []
2665
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   267
          |> abstract_over_list (rev qs)
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   268
      in
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   269
        RCInfo {RIvs=rcfix, rcarg=rcarg, CCas=rcassm, llRI=llRI, h_assum=h_assum}
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   270
      end
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   271
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   272
    val RC_infos = map2 mk_call_info RCs RIntro_thms
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   273
  in
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   274
    ClauseInfo {no=no, cdata=cdata, qglr=qglr, lGI=lGI, RCs=RC_infos,
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   275
      tree=tree}
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   276
  end
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   277
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   278
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   279
fun store_compat_thms 0 thms = []
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   280
  | store_compat_thms n thms =
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   281
  let
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   282
    val (thms1, thms2) = chop n thms
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   283
  in
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   284
    (thms1 :: store_compat_thms (n - 1) thms2)
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   285
  end
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   286
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   287
(* expects i <= j *)
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   288
fun lookup_compat_thm i j cts =
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   289
  nth (nth cts (i - 1)) (j - i)
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   290
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   291
(* nominal *)
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   292
(* Returns "Gsi, Gsj, lhs_i = lhs_j |-- rhs_j_f = rhs_i_f" *)
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   293
(* if j < i, then turn around *)
3239
67370521c09c updated for Isabelle 2015
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3229
diff changeset
   294
fun get_compat_thm ctxt cts eqvtsi eqvtsj invsi invsj i j ctxi ctxj =
2665
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   295
  let
2789
32979078bfe9 functions involving if and case do not throw exceptions anymore; but eqvt_at assumption has now a precondition
Christian Urban <urbanc@in.tum.de>
parents: 2788
diff changeset
   296
    val ClauseContext {cqs=cqsi,ags=agsi,lhs=lhsi,case_hyp=case_hypi,...} = ctxi
32979078bfe9 functions involving if and case do not throw exceptions anymore; but eqvt_at assumption has now a precondition
Christian Urban <urbanc@in.tum.de>
parents: 2788
diff changeset
   297
    val ClauseContext {cqs=cqsj,ags=agsj,lhs=lhsj,case_hyp=case_hypj,...} = ctxj
2665
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   298
3239
67370521c09c updated for Isabelle 2015
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3229
diff changeset
   299
    val lhsi_eq_lhsj = Thm.cterm_of ctxt (HOLogic.mk_Trueprop (mk_eq (lhsi, lhsj)))
2848
da7e6655cd4c fixed the problem when giving a complex default-term; the fundef lemmas in Nominal_Base were not general enough
Christian Urban <urbanc@in.tum.de>
parents: 2822
diff changeset
   300
2665
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   301
  in if j < i then
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   302
    let
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   303
      val compat = lookup_compat_thm j i cts
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   304
    in
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   305
      compat         (* "!!qj qi. Gsj => Gsi => lhsj = lhsi ==> rhsj = rhsi" *)
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   306
      |> fold Thm.forall_elim (cqsj @ cqsi) (* "Gsj => Gsi => lhsj = lhsi ==> rhsj = rhsi" *)
2789
32979078bfe9 functions involving if and case do not throw exceptions anymore; but eqvt_at assumption has now a precondition
Christian Urban <urbanc@in.tum.de>
parents: 2788
diff changeset
   307
      |> fold Thm.elim_implies eqvtsj (* nominal *)
2862
47063163f333 added eqvt_at and invariant for boths sides of the equations
Christian Urban <urbanc@in.tum.de>
parents: 2848
diff changeset
   308
      |> fold Thm.elim_implies eqvtsi (* nominal *)
2821
c7d4bd9e89e0 fixed problem with earlier commit about nominal_function_common; added facility for specifying an invariant - added a definition of frees_set which need a finiteness invariant
Christian Urban <urbanc@in.tum.de>
parents: 2819
diff changeset
   309
      |> fold Thm.elim_implies invsj (* nominal *)
2862
47063163f333 added eqvt_at and invariant for boths sides of the equations
Christian Urban <urbanc@in.tum.de>
parents: 2848
diff changeset
   310
      |> fold Thm.elim_implies invsi (* nominal *)
2665
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   311
      |> fold Thm.elim_implies agsj
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   312
      |> fold Thm.elim_implies agsi
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   313
      |> Thm.elim_implies ((Thm.assume lhsi_eq_lhsj) RS sym) (* "Gsj, Gsi, lhsi = lhsj |-- rhsj = rhsi" *)
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   314
    end
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   315
    else
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   316
    let
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   317
      val compat = lookup_compat_thm i j cts
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   318
    in
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   319
      compat        (* "!!qi qj. Gsi => Gsj => lhsi = lhsj ==> rhsi = rhsj" *)
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   320
      |> fold Thm.forall_elim (cqsi @ cqsj) (* "Gsi => Gsj => lhsi = lhsj ==> rhsi = rhsj" *)
2789
32979078bfe9 functions involving if and case do not throw exceptions anymore; but eqvt_at assumption has now a precondition
Christian Urban <urbanc@in.tum.de>
parents: 2788
diff changeset
   321
      |> fold Thm.elim_implies eqvtsi  (* nominal *)
2862
47063163f333 added eqvt_at and invariant for boths sides of the equations
Christian Urban <urbanc@in.tum.de>
parents: 2848
diff changeset
   322
      |> fold Thm.elim_implies eqvtsj  (* nominal *)
2821
c7d4bd9e89e0 fixed problem with earlier commit about nominal_function_common; added facility for specifying an invariant - added a definition of frees_set which need a finiteness invariant
Christian Urban <urbanc@in.tum.de>
parents: 2819
diff changeset
   323
      |> fold Thm.elim_implies invsi  (* nominal *)
2862
47063163f333 added eqvt_at and invariant for boths sides of the equations
Christian Urban <urbanc@in.tum.de>
parents: 2848
diff changeset
   324
      |> fold Thm.elim_implies invsj  (* nominal *)
2665
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   325
      |> fold Thm.elim_implies agsi
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   326
      |> fold Thm.elim_implies agsj
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   327
      |> Thm.elim_implies (Thm.assume lhsi_eq_lhsj)
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   328
      |> (fn thm => thm RS sym) (* "Gsi, Gsj, lhsi = lhsj |-- rhsj = rhsi" *)
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   329
    end
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   330
  end
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   331
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   332
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   333
(* Generates the replacement lemma in fully quantified form. *)
3239
67370521c09c updated for Isabelle 2015
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3229
diff changeset
   334
fun mk_replacement_lemma ctxt h ih_elim clause =
2665
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   335
  let
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   336
    val ClauseInfo {cdata=ClauseContext {qs, lhs, cqs, ags, case_hyp, ...},
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   337
      RCs, tree, ...} = clause
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   338
    local open Conv in
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   339
      val ih_conv = arg1_conv o arg_conv o arg_conv
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   340
    end
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   341
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   342
    val ih_elim_case =
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   343
      Conv.fconv_rule (ih_conv (K (case_hyp RS eq_reflection))) ih_elim
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   344
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   345
    val Ris = map (fn RCInfo {llRI, ...} => llRI) RCs
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   346
    val h_assums = map (fn RCInfo {h_assum, ...} =>
3239
67370521c09c updated for Isabelle 2015
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3229
diff changeset
   347
      Thm.assume (Thm.cterm_of ctxt (subst_bounds (rev qs, h_assum)))) RCs
2665
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   348
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   349
    val (eql, _) =
3239
67370521c09c updated for Isabelle 2015
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3229
diff changeset
   350
      Function_Context_Tree.rewrite_by_tree ctxt
3218
89158f401b07 updated to simplifier changes
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3200
diff changeset
   351
        h ih_elim_case (Ris ~~ h_assums) tree
2665
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   352
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   353
    val replace_lemma = (eql RS meta_eq_to_obj_eq)
3239
67370521c09c updated for Isabelle 2015
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3229
diff changeset
   354
      |> Thm.implies_intr (Thm.cprop_of case_hyp)
67370521c09c updated for Isabelle 2015
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3229
diff changeset
   355
      |> fold_rev (Thm.implies_intr o Thm.cprop_of) h_assums
67370521c09c updated for Isabelle 2015
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3229
diff changeset
   356
      |> fold_rev (Thm.implies_intr o Thm.cprop_of) ags
2665
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   357
      |> fold_rev Thm.forall_intr cqs
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   358
      |> Thm.close_derivation
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   359
  in
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   360
    replace_lemma
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   361
  end
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   362
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   363
(* nominal *)
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   364
(* Generates the eqvt lemmas for each clause *) 
3239
67370521c09c updated for Isabelle 2015
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3229
diff changeset
   365
fun mk_eqvt_lemma ctxt ih_eqvt clause =
2665
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   366
  let
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   367
    val ClauseInfo {cdata=ClauseContext {cqs, ags, case_hyp, ...}, RCs, ...} = clause
2789
32979078bfe9 functions involving if and case do not throw exceptions anymore; but eqvt_at assumption has now a precondition
Christian Urban <urbanc@in.tum.de>
parents: 2788
diff changeset
   368
     
2665
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   369
    local open Conv in
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   370
      val ih_conv = arg1_conv o arg_conv o arg_conv
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   371
    end
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   372
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   373
    val ih_eqvt_case =
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   374
      Conv.fconv_rule (ih_conv (K (case_hyp RS eq_reflection))) ih_eqvt
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   375
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   376
    fun prep_eqvt (RCInfo {llRI, RIvs, CCas, ...}) = 
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   377
        (llRI RS ih_eqvt_case)
3239
67370521c09c updated for Isabelle 2015
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3229
diff changeset
   378
        |> fold_rev (Thm.implies_intr o Thm.cprop_of) CCas
67370521c09c updated for Isabelle 2015
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3229
diff changeset
   379
        |> fold_rev (Thm.forall_intr o Thm.cterm_of ctxt o Free) RIvs 
2665
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   380
  in
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   381
    map prep_eqvt RCs
3239
67370521c09c updated for Isabelle 2015
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3229
diff changeset
   382
    |> map (fold_rev (Thm.implies_intr o Thm.cprop_of) ags)
67370521c09c updated for Isabelle 2015
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3229
diff changeset
   383
    |> map (Thm.implies_intr (Thm.cprop_of case_hyp))
2665
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   384
    |> map (fold_rev Thm.forall_intr cqs)
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   385
    |> map (Thm.close_derivation) 
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   386
  end
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   387
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   388
(* nominal *)
3239
67370521c09c updated for Isabelle 2015
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3229
diff changeset
   389
fun mk_invariant_lemma ctxt ih_inv clause =
2821
c7d4bd9e89e0 fixed problem with earlier commit about nominal_function_common; added facility for specifying an invariant - added a definition of frees_set which need a finiteness invariant
Christian Urban <urbanc@in.tum.de>
parents: 2819
diff changeset
   390
  let
c7d4bd9e89e0 fixed problem with earlier commit about nominal_function_common; added facility for specifying an invariant - added a definition of frees_set which need a finiteness invariant
Christian Urban <urbanc@in.tum.de>
parents: 2819
diff changeset
   391
    val ClauseInfo {cdata=ClauseContext {cqs, ags, case_hyp, ...}, RCs, ...} = clause
c7d4bd9e89e0 fixed problem with earlier commit about nominal_function_common; added facility for specifying an invariant - added a definition of frees_set which need a finiteness invariant
Christian Urban <urbanc@in.tum.de>
parents: 2819
diff changeset
   392
     
c7d4bd9e89e0 fixed problem with earlier commit about nominal_function_common; added facility for specifying an invariant - added a definition of frees_set which need a finiteness invariant
Christian Urban <urbanc@in.tum.de>
parents: 2819
diff changeset
   393
    local open Conv in
c7d4bd9e89e0 fixed problem with earlier commit about nominal_function_common; added facility for specifying an invariant - added a definition of frees_set which need a finiteness invariant
Christian Urban <urbanc@in.tum.de>
parents: 2819
diff changeset
   394
      val ih_conv = arg1_conv o arg_conv o arg_conv
c7d4bd9e89e0 fixed problem with earlier commit about nominal_function_common; added facility for specifying an invariant - added a definition of frees_set which need a finiteness invariant
Christian Urban <urbanc@in.tum.de>
parents: 2819
diff changeset
   395
    end
c7d4bd9e89e0 fixed problem with earlier commit about nominal_function_common; added facility for specifying an invariant - added a definition of frees_set which need a finiteness invariant
Christian Urban <urbanc@in.tum.de>
parents: 2819
diff changeset
   396
c7d4bd9e89e0 fixed problem with earlier commit about nominal_function_common; added facility for specifying an invariant - added a definition of frees_set which need a finiteness invariant
Christian Urban <urbanc@in.tum.de>
parents: 2819
diff changeset
   397
    val ih_inv_case =
c7d4bd9e89e0 fixed problem with earlier commit about nominal_function_common; added facility for specifying an invariant - added a definition of frees_set which need a finiteness invariant
Christian Urban <urbanc@in.tum.de>
parents: 2819
diff changeset
   398
      Conv.fconv_rule (ih_conv (K (case_hyp RS eq_reflection))) ih_inv
c7d4bd9e89e0 fixed problem with earlier commit about nominal_function_common; added facility for specifying an invariant - added a definition of frees_set which need a finiteness invariant
Christian Urban <urbanc@in.tum.de>
parents: 2819
diff changeset
   399
c7d4bd9e89e0 fixed problem with earlier commit about nominal_function_common; added facility for specifying an invariant - added a definition of frees_set which need a finiteness invariant
Christian Urban <urbanc@in.tum.de>
parents: 2819
diff changeset
   400
    fun prep_inv (RCInfo {llRI, RIvs, CCas, ...}) = 
c7d4bd9e89e0 fixed problem with earlier commit about nominal_function_common; added facility for specifying an invariant - added a definition of frees_set which need a finiteness invariant
Christian Urban <urbanc@in.tum.de>
parents: 2819
diff changeset
   401
        (llRI RS ih_inv_case)
3239
67370521c09c updated for Isabelle 2015
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3229
diff changeset
   402
        |> fold_rev (Thm.implies_intr o Thm.cprop_of) CCas
67370521c09c updated for Isabelle 2015
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3229
diff changeset
   403
        |> fold_rev (Thm.forall_intr o Thm.cterm_of ctxt o Free) RIvs 
2821
c7d4bd9e89e0 fixed problem with earlier commit about nominal_function_common; added facility for specifying an invariant - added a definition of frees_set which need a finiteness invariant
Christian Urban <urbanc@in.tum.de>
parents: 2819
diff changeset
   404
  in
c7d4bd9e89e0 fixed problem with earlier commit about nominal_function_common; added facility for specifying an invariant - added a definition of frees_set which need a finiteness invariant
Christian Urban <urbanc@in.tum.de>
parents: 2819
diff changeset
   405
    map prep_inv RCs
3239
67370521c09c updated for Isabelle 2015
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3229
diff changeset
   406
    |> map (fold_rev (Thm.implies_intr o Thm.cprop_of) ags)
67370521c09c updated for Isabelle 2015
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3229
diff changeset
   407
    |> map (Thm.implies_intr (Thm.cprop_of case_hyp))
2821
c7d4bd9e89e0 fixed problem with earlier commit about nominal_function_common; added facility for specifying an invariant - added a definition of frees_set which need a finiteness invariant
Christian Urban <urbanc@in.tum.de>
parents: 2819
diff changeset
   408
    |> map (fold_rev Thm.forall_intr cqs)
c7d4bd9e89e0 fixed problem with earlier commit about nominal_function_common; added facility for specifying an invariant - added a definition of frees_set which need a finiteness invariant
Christian Urban <urbanc@in.tum.de>
parents: 2819
diff changeset
   409
    |> map (Thm.close_derivation) 
c7d4bd9e89e0 fixed problem with earlier commit about nominal_function_common; added facility for specifying an invariant - added a definition of frees_set which need a finiteness invariant
Christian Urban <urbanc@in.tum.de>
parents: 2819
diff changeset
   410
  end
c7d4bd9e89e0 fixed problem with earlier commit about nominal_function_common; added facility for specifying an invariant - added a definition of frees_set which need a finiteness invariant
Christian Urban <urbanc@in.tum.de>
parents: 2819
diff changeset
   411
c7d4bd9e89e0 fixed problem with earlier commit about nominal_function_common; added facility for specifying an invariant - added a definition of frees_set which need a finiteness invariant
Christian Urban <urbanc@in.tum.de>
parents: 2819
diff changeset
   412
(* nominal *)
3239
67370521c09c updated for Isabelle 2015
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3229
diff changeset
   413
fun mk_uniqueness_clause ctxt globals compat_store eqvts invs clausei clausej RLj =
2665
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   414
  let
3239
67370521c09c updated for Isabelle 2015
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3229
diff changeset
   415
    val thy = Proof_Context.theory_of ctxt
67370521c09c updated for Isabelle 2015
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3229
diff changeset
   416
2665
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   417
    val Globals {h, y, x, fvar, ...} = globals
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   418
    val ClauseInfo {no=i, cdata=cctxi as ClauseContext {ctxt=ctxti, lhs=lhsi, case_hyp, cqs = cqsi, 
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   419
      ags = agsi, ...}, ...} = clausei
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   420
    val ClauseInfo {no=j, qglr=cdescj, RCs=RCsj, ...} = clausej
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   421
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   422
    val cctxj as ClauseContext {ags = agsj', lhs = lhsj', rhs = rhsj', qs = qsj', cqs = cqsj', ...} =
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   423
      mk_clause_context x ctxti cdescj 
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   424
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   425
    val rhsj'h = Pattern.rewrite_term thy [(fvar,h)] [] rhsj'
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   426
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   427
    val Ghsj' = map 
3239
67370521c09c updated for Isabelle 2015
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3229
diff changeset
   428
      (fn RCInfo {h_assum, ...} =>
67370521c09c updated for Isabelle 2015
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3229
diff changeset
   429
        Thm.assume (Thm.cterm_of ctxt (subst_bounds (rev qsj', h_assum)))) RCsj
2665
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   430
3239
67370521c09c updated for Isabelle 2015
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3229
diff changeset
   431
    val y_eq_rhsj'h = Thm.assume (Thm.cterm_of ctxt (HOLogic.mk_Trueprop (mk_eq (y, rhsj'h))))
67370521c09c updated for Isabelle 2015
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3229
diff changeset
   432
    val lhsi_eq_lhsj' = Thm.assume (Thm.cterm_of ctxt (HOLogic.mk_Trueprop (mk_eq (lhsi, lhsj'))))
2665
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   433
       (* lhs_i = lhs_j' |-- lhs_i = lhs_j' *) 
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   434
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   435
    val case_hypj' = trans OF [case_hyp, lhsi_eq_lhsj']
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   436
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   437
    val RLj_import = RLj
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   438
      |> fold Thm.forall_elim cqsj'
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   439
      |> fold Thm.elim_implies agsj'
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   440
      |> fold Thm.elim_implies Ghsj'
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   441
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   442
    val eqvtsi = nth eqvts (i - 1)
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   443
      |> map (fold Thm.forall_elim cqsi)
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   444
      |> map (fold Thm.elim_implies [case_hyp])
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   445
      |> map (fold Thm.elim_implies agsi)
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   446
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   447
    val eqvtsj = nth eqvts (j - 1)
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   448
      |> map (fold Thm.forall_elim cqsj')
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   449
      |> map (fold Thm.elim_implies [case_hypj'])
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   450
      |> map (fold Thm.elim_implies agsj')
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   451
2821
c7d4bd9e89e0 fixed problem with earlier commit about nominal_function_common; added facility for specifying an invariant - added a definition of frees_set which need a finiteness invariant
Christian Urban <urbanc@in.tum.de>
parents: 2819
diff changeset
   452
    val invsi = nth invs (i - 1)
c7d4bd9e89e0 fixed problem with earlier commit about nominal_function_common; added facility for specifying an invariant - added a definition of frees_set which need a finiteness invariant
Christian Urban <urbanc@in.tum.de>
parents: 2819
diff changeset
   453
      |> map (fold Thm.forall_elim cqsi)
c7d4bd9e89e0 fixed problem with earlier commit about nominal_function_common; added facility for specifying an invariant - added a definition of frees_set which need a finiteness invariant
Christian Urban <urbanc@in.tum.de>
parents: 2819
diff changeset
   454
      |> map (fold Thm.elim_implies [case_hyp])
c7d4bd9e89e0 fixed problem with earlier commit about nominal_function_common; added facility for specifying an invariant - added a definition of frees_set which need a finiteness invariant
Christian Urban <urbanc@in.tum.de>
parents: 2819
diff changeset
   455
      |> map (fold Thm.elim_implies agsi)
c7d4bd9e89e0 fixed problem with earlier commit about nominal_function_common; added facility for specifying an invariant - added a definition of frees_set which need a finiteness invariant
Christian Urban <urbanc@in.tum.de>
parents: 2819
diff changeset
   456
c7d4bd9e89e0 fixed problem with earlier commit about nominal_function_common; added facility for specifying an invariant - added a definition of frees_set which need a finiteness invariant
Christian Urban <urbanc@in.tum.de>
parents: 2819
diff changeset
   457
    val invsj = nth invs (j - 1)
c7d4bd9e89e0 fixed problem with earlier commit about nominal_function_common; added facility for specifying an invariant - added a definition of frees_set which need a finiteness invariant
Christian Urban <urbanc@in.tum.de>
parents: 2819
diff changeset
   458
      |> map (fold Thm.forall_elim cqsj')
c7d4bd9e89e0 fixed problem with earlier commit about nominal_function_common; added facility for specifying an invariant - added a definition of frees_set which need a finiteness invariant
Christian Urban <urbanc@in.tum.de>
parents: 2819
diff changeset
   459
      |> map (fold Thm.elim_implies [case_hypj'])
c7d4bd9e89e0 fixed problem with earlier commit about nominal_function_common; added facility for specifying an invariant - added a definition of frees_set which need a finiteness invariant
Christian Urban <urbanc@in.tum.de>
parents: 2819
diff changeset
   460
      |> map (fold Thm.elim_implies agsj')
c7d4bd9e89e0 fixed problem with earlier commit about nominal_function_common; added facility for specifying an invariant - added a definition of frees_set which need a finiteness invariant
Christian Urban <urbanc@in.tum.de>
parents: 2819
diff changeset
   461
3239
67370521c09c updated for Isabelle 2015
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3229
diff changeset
   462
    val compat = get_compat_thm ctxt compat_store eqvtsi eqvtsj invsi invsj i j cctxi cctxj
2665
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   463
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   464
  in
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   465
    (trans OF [case_hyp, lhsi_eq_lhsj']) (* lhs_i = lhs_j' |-- x = lhs_j' *)
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   466
    |> Thm.implies_elim RLj_import
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   467
      (* Rj1' ... Rjk', lhs_i = lhs_j' |-- rhs_j'_h = rhs_j'_f *)
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   468
    |> (fn it => trans OF [it, compat])
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   469
      (* lhs_i = lhs_j', Gj', Rj1' ... Rjk' |-- rhs_j'_h = rhs_i_f *)
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   470
    |> (fn it => trans OF [y_eq_rhsj'h, it])
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   471
      (* lhs_i = lhs_j', Gj', Rj1' ... Rjk', y = rhs_j_h' |-- y = rhs_i_f *)
3239
67370521c09c updated for Isabelle 2015
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3229
diff changeset
   472
    |> fold_rev (Thm.implies_intr o Thm.cprop_of) Ghsj'
67370521c09c updated for Isabelle 2015
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3229
diff changeset
   473
    |> fold_rev (Thm.implies_intr o Thm.cprop_of) agsj'
2665
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   474
      (* lhs_i = lhs_j' , y = rhs_j_h' |-- Gj', Rj1'...Rjk' ==> y = rhs_i_f *)
3239
67370521c09c updated for Isabelle 2015
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3229
diff changeset
   475
    |> Thm.implies_intr (Thm.cprop_of y_eq_rhsj'h)
67370521c09c updated for Isabelle 2015
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3229
diff changeset
   476
    |> Thm.implies_intr (Thm.cprop_of lhsi_eq_lhsj')
67370521c09c updated for Isabelle 2015
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3229
diff changeset
   477
    |> fold_rev Thm.forall_intr (Thm.cterm_of ctxt h :: cqsj')
2665
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   478
  end
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   479
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   480
(* nominal *)
3239
67370521c09c updated for Isabelle 2015
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3229
diff changeset
   481
fun mk_uniqueness_case ctxt
67370521c09c updated for Isabelle 2015
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3229
diff changeset
   482
    globals G f ihyp ih_intro G_cases compat_store clauses replems eqvtlems invlems clausei =
2665
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   483
  let
3239
67370521c09c updated for Isabelle 2015
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3229
diff changeset
   484
    val thy = Proof_Context.theory_of ctxt
67370521c09c updated for Isabelle 2015
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3229
diff changeset
   485
2665
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   486
    val Globals {x, y, ranT, fvar, ...} = globals
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   487
    val ClauseInfo {cdata = ClauseContext {lhs, rhs, cqs, ags, case_hyp, ...}, lGI, RCs, ...} = clausei
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   488
    val rhsC = Pattern.rewrite_term thy [(fvar, f)] [] rhs
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   489
3218
89158f401b07 updated to simplifier changes
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3200
diff changeset
   490
    val ih_intro_case =
3239
67370521c09c updated for Isabelle 2015
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3229
diff changeset
   491
      full_simplify (put_simpset HOL_basic_ss ctxt addsimps [case_hyp])
3227
35bb5b013f0e updated with current AFP version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3220
diff changeset
   492
        ih_intro
2665
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   493
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   494
    fun prep_RC (RCInfo {llRI, RIvs, CCas, ...}) = 
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   495
        (llRI RS ih_intro_case)
3239
67370521c09c updated for Isabelle 2015
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3229
diff changeset
   496
        |> fold_rev (Thm.implies_intr o Thm.cprop_of) CCas
67370521c09c updated for Isabelle 2015
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3229
diff changeset
   497
        |> fold_rev (Thm.forall_intr o Thm.cterm_of ctxt o Free) RIvs
2789
32979078bfe9 functions involving if and case do not throw exceptions anymore; but eqvt_at assumption has now a precondition
Christian Urban <urbanc@in.tum.de>
parents: 2788
diff changeset
   498
 
2665
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   499
    val existence = fold (curry op COMP o prep_RC) RCs lGI
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   500
3239
67370521c09c updated for Isabelle 2015
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3229
diff changeset
   501
    val P = Thm.cterm_of ctxt (mk_eq (y, rhsC))
67370521c09c updated for Isabelle 2015
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3229
diff changeset
   502
    val G_lhs_y = Thm.assume (Thm.cterm_of ctxt (HOLogic.mk_Trueprop (G $ lhs $ y)))
2665
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   503
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   504
    val unique_clauses =
3239
67370521c09c updated for Isabelle 2015
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3229
diff changeset
   505
      map2 (mk_uniqueness_clause ctxt globals compat_store eqvtlems invlems clausei) clauses replems
2665
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   506
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   507
    fun elim_implies_eta A AB =
3239
67370521c09c updated for Isabelle 2015
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3229
diff changeset
   508
      Thm.bicompose NONE {flatten = false, match = true, incremented = false} (false, A, 0) 1 AB
3219
e5d9b6bca88c updated to new Isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3218
diff changeset
   509
      |> Seq.list_of |> the_single
e5d9b6bca88c updated to new Isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3218
diff changeset
   510
2665
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   511
    val uniqueness = G_cases
3239
67370521c09c updated for Isabelle 2015
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3229
diff changeset
   512
      |> Thm.forall_elim (Thm.cterm_of ctxt lhs)
67370521c09c updated for Isabelle 2015
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3229
diff changeset
   513
      |> Thm.forall_elim (Thm.cterm_of ctxt y)
2665
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   514
      |> Thm.forall_elim P
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   515
      |> Thm.elim_implies G_lhs_y
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   516
      |> fold elim_implies_eta unique_clauses
3239
67370521c09c updated for Isabelle 2015
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3229
diff changeset
   517
      |> Thm.implies_intr (Thm.cprop_of G_lhs_y)
67370521c09c updated for Isabelle 2015
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3229
diff changeset
   518
      |> Thm.forall_intr (Thm.cterm_of ctxt y)
2665
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   519
3239
67370521c09c updated for Isabelle 2015
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3229
diff changeset
   520
    val P2 = Thm.cterm_of ctxt (lambda y (G $ lhs $ y)) (* P2 y := (lhs, y): G *)
2665
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   521
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   522
    val exactly_one =
3244
a44479bde681 fixed a problem with two example theories
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3239
diff changeset
   523
      ex1I |> Thm.instantiate' [SOME (Thm.ctyp_of ctxt ranT)] [SOME P2, SOME (Thm.cterm_of ctxt rhsC)]
2665
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   524
      |> curry (op COMP) existence
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   525
      |> curry (op COMP) uniqueness
3239
67370521c09c updated for Isabelle 2015
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3229
diff changeset
   526
      |> simplify (put_simpset HOL_basic_ss ctxt addsimps [case_hyp RS sym])
67370521c09c updated for Isabelle 2015
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3229
diff changeset
   527
      |> Thm.implies_intr (Thm.cprop_of case_hyp)
67370521c09c updated for Isabelle 2015
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3229
diff changeset
   528
      |> fold_rev (Thm.implies_intr o Thm.cprop_of) ags
2665
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   529
      |> fold_rev Thm.forall_intr cqs
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   530
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   531
    val function_value =
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   532
      existence
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   533
      |> Thm.implies_intr ihyp
3239
67370521c09c updated for Isabelle 2015
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3229
diff changeset
   534
      |> Thm.implies_intr (Thm.cprop_of case_hyp)
67370521c09c updated for Isabelle 2015
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3229
diff changeset
   535
      |> Thm.forall_intr (Thm.cterm_of ctxt x)
67370521c09c updated for Isabelle 2015
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3229
diff changeset
   536
      |> Thm.forall_elim (Thm.cterm_of ctxt lhs)
2665
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   537
      |> curry (op RS) refl
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   538
  in
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   539
    (exactly_one, function_value)
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   540
  end
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   541
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   542
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   543
(* nominal *)
3239
67370521c09c updated for Isabelle 2015
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3229
diff changeset
   544
fun prove_stuff ctxt
67370521c09c updated for Isabelle 2015
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3229
diff changeset
   545
    globals G f R clauses complete compat compat_store G_elim G_eqvt invariant f_def =
2665
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   546
  let
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   547
    val Globals {h, domT, ranT, x, ...} = globals
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   548
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   549
    (* Inductive Hypothesis: !!z. (z,x):R ==> EX!y. (z,y):G *)
3108
61db5ad429bb updated to Isabelle 16 January
Christian Urban <urbanc@in.tum.de>
parents: 3045
diff changeset
   550
    val ihyp = Logic.all_const domT $ Abs ("z", domT,
2665
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   551
      Logic.mk_implies (HOLogic.mk_Trueprop (R $ Bound 0 $ x),
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   552
        HOLogic.mk_Trueprop (Const (@{const_name Ex1}, (ranT --> boolT) --> boolT) $
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   553
          Abs ("y", ranT, G $ Bound 1 $ Bound 0))))
3239
67370521c09c updated for Isabelle 2015
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3229
diff changeset
   554
      |> Thm.cterm_of ctxt
2665
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   555
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   556
    val ihyp_thm = Thm.assume ihyp |> Thm.forall_elim_vars 0
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   557
    val ih_intro = ihyp_thm RS (f_def RS ex1_implies_ex)
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   558
    val ih_elim = ihyp_thm RS (f_def RS ex1_implies_un)
3244
a44479bde681 fixed a problem with two example theories
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3239
diff changeset
   559
      |> Thm.instantiate' [] [NONE, SOME (Thm.cterm_of ctxt h)]
2665
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   560
    val ih_eqvt = ihyp_thm RS (G_eqvt RS (f_def RS @{thm fundef_ex1_eqvt_at}))
2821
c7d4bd9e89e0 fixed problem with earlier commit about nominal_function_common; added facility for specifying an invariant - added a definition of frees_set which need a finiteness invariant
Christian Urban <urbanc@in.tum.de>
parents: 2819
diff changeset
   561
    val ih_inv =  ihyp_thm RS (invariant COMP (f_def RS @{thm fundef_ex1_prop}))
c7d4bd9e89e0 fixed problem with earlier commit about nominal_function_common; added facility for specifying an invariant - added a definition of frees_set which need a finiteness invariant
Christian Urban <urbanc@in.tum.de>
parents: 2819
diff changeset
   562
2665
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   563
    val _ = trace_msg (K "Proving Replacement lemmas...")
3239
67370521c09c updated for Isabelle 2015
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3229
diff changeset
   564
    val repLemmas = map (mk_replacement_lemma ctxt h ih_elim) clauses
2665
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   565
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   566
    val _ = trace_msg (K "Proving Equivariance lemmas...")
3239
67370521c09c updated for Isabelle 2015
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3229
diff changeset
   567
    val eqvtLemmas = map (mk_eqvt_lemma ctxt ih_eqvt) clauses
2665
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   568
2821
c7d4bd9e89e0 fixed problem with earlier commit about nominal_function_common; added facility for specifying an invariant - added a definition of frees_set which need a finiteness invariant
Christian Urban <urbanc@in.tum.de>
parents: 2819
diff changeset
   569
    val _ = trace_msg (K "Proving Invariance lemmas...")
3239
67370521c09c updated for Isabelle 2015
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3229
diff changeset
   570
    val invLemmas = map (mk_invariant_lemma ctxt ih_inv) clauses
2821
c7d4bd9e89e0 fixed problem with earlier commit about nominal_function_common; added facility for specifying an invariant - added a definition of frees_set which need a finiteness invariant
Christian Urban <urbanc@in.tum.de>
parents: 2819
diff changeset
   571
2665
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   572
    val _ = trace_msg (K "Proving cases for unique existence...")
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   573
    val (ex1s, values) =
3239
67370521c09c updated for Isabelle 2015
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3229
diff changeset
   574
      split_list (map (mk_uniqueness_case ctxt globals G f 
2821
c7d4bd9e89e0 fixed problem with earlier commit about nominal_function_common; added facility for specifying an invariant - added a definition of frees_set which need a finiteness invariant
Christian Urban <urbanc@in.tum.de>
parents: 2819
diff changeset
   575
        ihyp ih_intro G_elim compat_store clauses repLemmas eqvtLemmas invLemmas) clauses)
2665
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   576
     
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   577
    val _ = trace_msg (K "Proving: Graph is a function")
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   578
    val graph_is_function = complete
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   579
      |> Thm.forall_elim_vars 0
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   580
      |> fold (curry op COMP) ex1s
3239
67370521c09c updated for Isabelle 2015
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3229
diff changeset
   581
      |> Thm.implies_intr ihyp
67370521c09c updated for Isabelle 2015
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3229
diff changeset
   582
      |> Thm.implies_intr (Thm.cterm_of ctxt (HOLogic.mk_Trueprop (mk_acc domT R $ x)))
67370521c09c updated for Isabelle 2015
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3229
diff changeset
   583
      |> Thm.forall_intr (Thm.cterm_of ctxt x)
3220
87dbeba4b25a updated to ne Isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3219
diff changeset
   584
      |> (fn it => Drule.compose (it, 2, acc_induct_rule)) (* "EX! y. (?x,y):G" *)
3239
67370521c09c updated for Isabelle 2015
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3229
diff changeset
   585
      |> (fn it =>
67370521c09c updated for Isabelle 2015
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3229
diff changeset
   586
          fold (Thm.forall_intr o Thm.cterm_of ctxt o Var) (Term.add_vars (Thm.prop_of it) []) it)
2665
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   587
2707
747ebf2f066d made eqvt-proof explicit in the function definitions
Christian Urban <urbanc@in.tum.de>
parents: 2677
diff changeset
   588
    val goalstate =  
3239
67370521c09c updated for Isabelle 2015
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3229
diff changeset
   589
      Conjunction.intr (Conjunction.intr (Conjunction.intr graph_is_function complete) invariant) G_eqvt
2665
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   590
      |> Thm.close_derivation
3220
87dbeba4b25a updated to ne Isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3219
diff changeset
   591
      |> Goal.protect 0
3239
67370521c09c updated for Isabelle 2015
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3229
diff changeset
   592
      |> fold_rev (Thm.implies_intr o Thm.cprop_of) compat
67370521c09c updated for Isabelle 2015
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3229
diff changeset
   593
      |> Thm.implies_intr (Thm.cprop_of complete)
67370521c09c updated for Isabelle 2015
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3229
diff changeset
   594
      |> Thm.implies_intr (Thm.cprop_of invariant)
67370521c09c updated for Isabelle 2015
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3229
diff changeset
   595
      |> Thm.implies_intr (Thm.cprop_of G_eqvt)
2665
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   596
  in
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   597
    (goalstate, values)
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   598
  end
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   599
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   600
(* wrapper -- restores quantifiers in rule specifications *)
2707
747ebf2f066d made eqvt-proof explicit in the function definitions
Christian Urban <urbanc@in.tum.de>
parents: 2677
diff changeset
   601
fun inductive_def (binding as ((R, T), _)) intrs lthy =
3197
25d11b449e92 definition of an auxiliary graph in nominal-primrec definitions
Christian Urban <urbanc@in.tum.de>
parents: 3108
diff changeset
   602
  let 
2665
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   603
    val ({intrs = intrs_gen, elims = [elim_gen], preds = [ Rdef ], induct, raw_induct, ...}, lthy) =
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   604
      lthy
3244
a44479bde681 fixed a problem with two example theories
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3239
diff changeset
   605
      |> Config.put Inductive.inductive_internals true
3239
67370521c09c updated for Isabelle 2015
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3229
diff changeset
   606
      |> Proof_Context.concealed
2665
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   607
      |> Inductive.add_inductive_i
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   608
          {quiet_mode = true,
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   609
            verbose = ! trace,
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   610
            alt_name = Binding.empty,
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   611
            coind = false,
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   612
            no_elim = false,
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   613
            no_ind = false,
3200
995d47b09ab4 removed fork_mono flag
Christian Urban <urbanc@in.tum.de>
parents: 3197
diff changeset
   614
            skip_mono = true}
2665
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   615
          [binding] (* relation *)
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   616
          [] (* no parameters *)
3245
017e33849f4d updated to Isabelle 2016-1
Christian Urban <urbanc@in.tum.de>
parents: 3244
diff changeset
   617
          (map (fn t => (Binding.empty_atts, t)) intrs) (* intro rules *)
2665
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   618
          [] (* no special monos *)
3239
67370521c09c updated for Isabelle 2015
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3229
diff changeset
   619
      ||> Proof_Context.restore_naming lthy
3244
a44479bde681 fixed a problem with two example theories
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3239
diff changeset
   620
      ||> Config.put Inductive.inductive_internals (Config.get lthy Inductive.inductive_internals)
2665
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   621
3239
67370521c09c updated for Isabelle 2015
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3229
diff changeset
   622
    val cert = Thm.cterm_of lthy
2665
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   623
    fun requantify orig_intro thm =
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   624
      let
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   625
        val (qs, t) = dest_all_all orig_intro
2781
542ff50555f5 updated to new Isabelle (> 9 May)
Christian Urban <urbanc@in.tum.de>
parents: 2745
diff changeset
   626
        val frees = Variable.add_frees lthy t [] |> remove (op =) (Binding.name_of R, T)
3239
67370521c09c updated for Isabelle 2015
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3229
diff changeset
   627
        val vars = Term.add_vars (Thm.prop_of thm) []
2665
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   628
        val varmap = AList.lookup (op =) (frees ~~ map fst vars)
2781
542ff50555f5 updated to new Isabelle (> 9 May)
Christian Urban <urbanc@in.tum.de>
parents: 2745
diff changeset
   629
           #> the_default ("",0)
2665
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   630
      in
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   631
        fold_rev (fn Free (n, T) =>
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   632
          forall_intr_rename (n, cert (Var (varmap (n, T), T)))) qs thm
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   633
      end
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   634
  in
2707
747ebf2f066d made eqvt-proof explicit in the function definitions
Christian Urban <urbanc@in.tum.de>
parents: 2677
diff changeset
   635
    ((Rdef, map2 requantify intrs intrs_gen, forall_intr_vars elim_gen, induct), lthy)
2665
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   636
  end
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   637
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   638
(* nominal *)
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   639
fun define_graph Gname fvar domT ranT clauses RCss lthy =
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   640
  let
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   641
    val GT = domT --> ranT --> boolT
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   642
    val (Gvar as (n, T)) = singleton (Variable.variant_frees lthy []) (Gname, GT)
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   643
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   644
    fun mk_GIntro (ClauseContext {qs, gs, lhs, rhs, ...}) RCs =
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   645
      let
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   646
        fun mk_h_assm (rcfix, rcassm, rcarg) =
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   647
          HOLogic.mk_Trueprop (Free Gvar $ rcarg $ (fvar $ rcarg))
3239
67370521c09c updated for Isabelle 2015
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3229
diff changeset
   648
          |> fold_rev (curry Logic.mk_implies o Thm.prop_of) rcassm
2665
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   649
          |> fold_rev (Logic.all o Free) rcfix
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   650
      in
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   651
        HOLogic.mk_Trueprop (Free Gvar $ lhs $ rhs)
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   652
        |> fold_rev (curry Logic.mk_implies o mk_h_assm) RCs
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   653
        |> fold_rev (curry Logic.mk_implies) gs
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   654
        |> fold_rev Logic.all (fvar :: qs)
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   655
      end
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   656
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   657
    val G_intros = map2 mk_GIntro clauses RCss
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   658
  in
2707
747ebf2f066d made eqvt-proof explicit in the function definitions
Christian Urban <urbanc@in.tum.de>
parents: 2677
diff changeset
   659
    inductive_def ((Binding.name n, T), NoSyn) G_intros lthy
2665
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   660
  end
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   661
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   662
fun define_function fdefname (fname, mixfix) domT ranT G default lthy =
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   663
  let
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   664
    val f_def =
3229
b52e8651591f updated to Isabelle changes
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3227
diff changeset
   665
      Abs ("x", domT, Const (@{const_name Fun_Def.THE_default}, ranT --> (ranT --> boolT) --> ranT) 
2665
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   666
        $ (default $ Bound 0) $ Abs ("y", ranT, G $ Bound 1 $ Bound 0))
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   667
      |> Syntax.check_term lthy
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   668
  in
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   669
    Local_Theory.define
3245
017e33849f4d updated to Isabelle 2016-1
Christian Urban <urbanc@in.tum.de>
parents: 3244
diff changeset
   670
      ((Binding.name (fname ^ "C"), mixfix),
3239
67370521c09c updated for Isabelle 2015
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3229
diff changeset
   671
        ((Binding.concealed (Binding.name fdefname), []), f_def)) lthy
2665
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   672
  end
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   673
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   674
(* nominal *)
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   675
fun define_recursion_relation Rname domT qglrs clauses RCss lthy =
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   676
  let
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   677
    val RT = domT --> domT --> boolT
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   678
    val (Rvar as (n, T)) = singleton (Variable.variant_frees lthy []) (Rname, RT)
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   679
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   680
    fun mk_RIntro (ClauseContext {qs, gs, lhs, ...}, (oqs, _, _, _)) (rcfix, rcassm, rcarg) =
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   681
      HOLogic.mk_Trueprop (Free Rvar $ rcarg $ lhs)
3239
67370521c09c updated for Isabelle 2015
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3229
diff changeset
   682
      |> fold_rev (curry Logic.mk_implies o Thm.prop_of) rcassm
2665
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   683
      |> fold_rev (curry Logic.mk_implies) gs
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   684
      |> fold_rev (Logic.all o Free) rcfix
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   685
      |> fold_rev mk_forall_rename (map fst oqs ~~ qs)
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   686
      (* "!!qs xs. CS ==> G => (r, lhs) : R" *)
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   687
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   688
    val R_intross = map2 (map o mk_RIntro) (clauses ~~ qglrs) RCss
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   689
2707
747ebf2f066d made eqvt-proof explicit in the function definitions
Christian Urban <urbanc@in.tum.de>
parents: 2677
diff changeset
   690
    val ((R, RIntro_thms, R_elim, _), lthy) =
747ebf2f066d made eqvt-proof explicit in the function definitions
Christian Urban <urbanc@in.tum.de>
parents: 2677
diff changeset
   691
      inductive_def ((Binding.name n, T), NoSyn) (flat R_intross) lthy
2665
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   692
  in
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   693
    ((R, Library.unflat R_intross RIntro_thms, R_elim), lthy)
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   694
  end
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   695
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   696
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   697
fun fix_globals domT ranT fvar ctxt =
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   698
  let
3239
67370521c09c updated for Isabelle 2015
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3229
diff changeset
   699
    val ([h, y, x, z, a, D, P, Pbool], ctxt') = Variable.variant_fixes
2665
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   700
      ["h_fd", "y_fd", "x_fd", "z_fd", "a_fd", "D_fd", "P_fd", "Pb_fd"] ctxt
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   701
  in
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   702
    (Globals {h = Free (h, domT --> ranT),
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   703
      y = Free (y, ranT),
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   704
      x = Free (x, domT),
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   705
      z = Free (z, domT),
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   706
      a = Free (a, domT),
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   707
      D = Free (D, domT --> boolT),
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   708
      P = Free (P, domT --> boolT),
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   709
      Pbool = Free (Pbool, boolT),
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   710
      fvar = fvar,
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   711
      domT = domT,
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   712
      ranT = ranT},
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   713
    ctxt')
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   714
  end
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   715
3239
67370521c09c updated for Isabelle 2015
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3229
diff changeset
   716
fun inst_RC ctxt fvar f (rcfix, rcassm, rcarg) =
2665
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   717
  let
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   718
    fun inst_term t = subst_bound(f, abstract_over (fvar, t))
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   719
  in
3239
67370521c09c updated for Isabelle 2015
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3229
diff changeset
   720
    (rcfix, map (Thm.assume o Thm.cterm_of ctxt o inst_term o Thm.prop_of) rcassm, inst_term rcarg)
2665
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   721
  end
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   722
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   723
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   724
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   725
(**********************************************************
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   726
 *                   PROVING THE RULES
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   727
 **********************************************************)
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   728
3239
67370521c09c updated for Isabelle 2015
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3229
diff changeset
   729
fun mk_psimps ctxt globals R clauses valthms f_iff graph_is_function =
2665
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   730
  let
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   731
    val Globals {domT, z, ...} = globals
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   732
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   733
    fun mk_psimp (ClauseInfo {qglr = (oqs, _, _, _), cdata = ClauseContext {cqs, lhs, ags, ...}, ...}) valthm =
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   734
      let
3239
67370521c09c updated for Isabelle 2015
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3229
diff changeset
   735
        val lhs_acc = Thm.cterm_of ctxt (HOLogic.mk_Trueprop (mk_acc domT R $ lhs)) (* "acc R lhs" *)
67370521c09c updated for Isabelle 2015
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3229
diff changeset
   736
        val z_smaller = Thm.cterm_of ctxt (HOLogic.mk_Trueprop (R $ z $ lhs)) (* "R z lhs" *)
2665
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   737
      in
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   738
        ((Thm.assume z_smaller) RS ((Thm.assume lhs_acc) RS acc_downward))
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   739
        |> (fn it => it COMP graph_is_function)
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   740
        |> Thm.implies_intr z_smaller
3239
67370521c09c updated for Isabelle 2015
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3229
diff changeset
   741
        |> Thm.forall_intr (Thm.cterm_of ctxt z)
2665
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   742
        |> (fn it => it COMP valthm)
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   743
        |> Thm.implies_intr lhs_acc
3239
67370521c09c updated for Isabelle 2015
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3229
diff changeset
   744
        |> asm_simplify (put_simpset HOL_basic_ss ctxt addsimps [f_iff])
67370521c09c updated for Isabelle 2015
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3229
diff changeset
   745
        |> fold_rev (Thm.implies_intr o Thm.cprop_of) ags
2665
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   746
        |> fold_rev forall_intr_rename (map fst oqs ~~ cqs)
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   747
      end
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   748
  in
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   749
    map2 mk_psimp clauses valthms
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   750
  end
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   751
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   752
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   753
(** Induction rule **)
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   754
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   755
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   756
val acc_subset_induct = @{thm predicate1I} RS @{thm accp_subset_induct}
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   757
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   758
3239
67370521c09c updated for Isabelle 2015
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3229
diff changeset
   759
fun mk_partial_induct_rule ctxt globals R complete_thm clauses =
2665
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   760
  let
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   761
    val Globals {domT, x, z, a, P, D, ...} = globals
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   762
    val acc_R = mk_acc domT R
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   763
3239
67370521c09c updated for Isabelle 2015
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3229
diff changeset
   764
    val x_D = Thm.assume (Thm.cterm_of ctxt (HOLogic.mk_Trueprop (D $ x)))
67370521c09c updated for Isabelle 2015
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3229
diff changeset
   765
    val a_D = Thm.cterm_of ctxt (HOLogic.mk_Trueprop (D $ a))
2665
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   766
3239
67370521c09c updated for Isabelle 2015
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3229
diff changeset
   767
    val D_subset = Thm.cterm_of ctxt (Logic.all x
2665
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   768
      (Logic.mk_implies (HOLogic.mk_Trueprop (D $ x), HOLogic.mk_Trueprop (acc_R $ x))))
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   769
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   770
    val D_dcl = (* "!!x z. [| x: D; (z,x):R |] ==> z:D" *)
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   771
      Logic.all x (Logic.all z (Logic.mk_implies (HOLogic.mk_Trueprop (D $ x),
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   772
        Logic.mk_implies (HOLogic.mk_Trueprop (R $ z $ x),
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   773
          HOLogic.mk_Trueprop (D $ z)))))
3239
67370521c09c updated for Isabelle 2015
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3229
diff changeset
   774
      |> Thm.cterm_of ctxt
2665
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   775
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   776
    (* Inductive Hypothesis: !!z. (z,x):R ==> P z *)
3108
61db5ad429bb updated to Isabelle 16 January
Christian Urban <urbanc@in.tum.de>
parents: 3045
diff changeset
   777
    val ihyp = Logic.all_const domT $ Abs ("z", domT,
2665
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   778
      Logic.mk_implies (HOLogic.mk_Trueprop (R $ Bound 0 $ x),
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   779
        HOLogic.mk_Trueprop (P $ Bound 0)))
3239
67370521c09c updated for Isabelle 2015
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3229
diff changeset
   780
      |> Thm.cterm_of ctxt
2665
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   781
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   782
    val aihyp = Thm.assume ihyp
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   783
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   784
    fun prove_case clause =
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   785
      let
3239
67370521c09c updated for Isabelle 2015
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3229
diff changeset
   786
        val ClauseInfo {cdata = ClauseContext {ctxt = ctxt', qs, cqs, ags, gs, lhs, case_hyp, ...},
2665
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   787
          RCs, qglr = (oqs, _, _, _), ...} = clause
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   788
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   789
        val case_hyp_conv = K (case_hyp RS eq_reflection)
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   790
        local open Conv in
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   791
          val lhs_D = fconv_rule (arg_conv (arg_conv (case_hyp_conv))) x_D
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   792
          val sih =
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   793
            fconv_rule (Conv.binder_conv
3239
67370521c09c updated for Isabelle 2015
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3229
diff changeset
   794
              (K (arg1_conv (arg_conv (arg_conv case_hyp_conv)))) ctxt') aihyp
2665
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   795
        end
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   796
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   797
        fun mk_Prec (RCInfo {llRI, RIvs, CCas, rcarg, ...}) = sih
3239
67370521c09c updated for Isabelle 2015
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3229
diff changeset
   798
          |> Thm.forall_elim (Thm.cterm_of ctxt rcarg)
2665
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   799
          |> Thm.elim_implies llRI
3239
67370521c09c updated for Isabelle 2015
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3229
diff changeset
   800
          |> fold_rev (Thm.implies_intr o Thm.cprop_of) CCas
67370521c09c updated for Isabelle 2015
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3229
diff changeset
   801
          |> fold_rev (Thm.forall_intr o Thm.cterm_of ctxt o Free) RIvs
2665
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   802
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   803
        val P_recs = map mk_Prec RCs   (*  [P rec1, P rec2, ... ]  *)
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   804
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   805
        val step = HOLogic.mk_Trueprop (P $ lhs)
3239
67370521c09c updated for Isabelle 2015
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3229
diff changeset
   806
          |> fold_rev (curry Logic.mk_implies o Thm.prop_of) P_recs
2665
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   807
          |> fold_rev (curry Logic.mk_implies) gs
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   808
          |> curry Logic.mk_implies (HOLogic.mk_Trueprop (D $ lhs))
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   809
          |> fold_rev mk_forall_rename (map fst oqs ~~ qs)
3239
67370521c09c updated for Isabelle 2015
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3229
diff changeset
   810
          |> Thm.cterm_of ctxt
2665
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   811
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   812
        val P_lhs = Thm.assume step
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   813
          |> fold Thm.forall_elim cqs
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   814
          |> Thm.elim_implies lhs_D
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   815
          |> fold Thm.elim_implies ags
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   816
          |> fold Thm.elim_implies P_recs
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   817
3239
67370521c09c updated for Isabelle 2015
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3229
diff changeset
   818
        val res =
67370521c09c updated for Isabelle 2015
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3229
diff changeset
   819
          Thm.cterm_of ctxt (HOLogic.mk_Trueprop (P $ x))
2665
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   820
          |> Conv.arg_conv (Conv.arg_conv case_hyp_conv)
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   821
          |> Thm.symmetric (* P lhs == P x *)
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   822
          |> (fn eql => Thm.equal_elim eql P_lhs) (* "P x" *)
3239
67370521c09c updated for Isabelle 2015
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3229
diff changeset
   823
          |> Thm.implies_intr (Thm.cprop_of case_hyp)
67370521c09c updated for Isabelle 2015
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3229
diff changeset
   824
          |> fold_rev (Thm.implies_intr o Thm.cprop_of) ags
2665
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   825
          |> fold_rev Thm.forall_intr cqs
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   826
      in
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   827
        (res, step)
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   828
      end
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   829
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   830
    val (cases, steps) = split_list (map prove_case clauses)
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   831
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   832
    val istep = complete_thm
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   833
      |> Thm.forall_elim_vars 0
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   834
      |> fold (curry op COMP) cases (*  P x  *)
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   835
      |> Thm.implies_intr ihyp
3239
67370521c09c updated for Isabelle 2015
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3229
diff changeset
   836
      |> Thm.implies_intr (Thm.cprop_of x_D)
67370521c09c updated for Isabelle 2015
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3229
diff changeset
   837
      |> Thm.forall_intr (Thm.cterm_of ctxt x)
2665
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   838
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   839
    val subset_induct_rule =
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   840
      acc_subset_induct
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   841
      |> (curry op COMP) (Thm.assume D_subset)
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   842
      |> (curry op COMP) (Thm.assume D_dcl)
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   843
      |> (curry op COMP) (Thm.assume a_D)
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   844
      |> (curry op COMP) istep
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   845
      |> fold_rev Thm.implies_intr steps
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   846
      |> Thm.implies_intr a_D
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   847
      |> Thm.implies_intr D_dcl
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   848
      |> Thm.implies_intr D_subset
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   849
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   850
    val simple_induct_rule =
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   851
      subset_induct_rule
3239
67370521c09c updated for Isabelle 2015
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3229
diff changeset
   852
      |> Thm.forall_intr (Thm.cterm_of ctxt D)
67370521c09c updated for Isabelle 2015
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3229
diff changeset
   853
      |> Thm.forall_elim (Thm.cterm_of ctxt acc_R)
3244
a44479bde681 fixed a problem with two example theories
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3239
diff changeset
   854
      |> assume_tac ctxt 1 |> Seq.hd
2665
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   855
      |> (curry op COMP) (acc_downward
3244
a44479bde681 fixed a problem with two example theories
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3239
diff changeset
   856
        |> (Thm.instantiate' [SOME (Thm.ctyp_of ctxt domT)]
3239
67370521c09c updated for Isabelle 2015
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3229
diff changeset
   857
             (map (SOME o Thm.cterm_of ctxt) [R, x, z]))
67370521c09c updated for Isabelle 2015
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3229
diff changeset
   858
        |> Thm.forall_intr (Thm.cterm_of ctxt z)
67370521c09c updated for Isabelle 2015
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3229
diff changeset
   859
        |> Thm.forall_intr (Thm.cterm_of ctxt x))
67370521c09c updated for Isabelle 2015
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3229
diff changeset
   860
      |> Thm.forall_intr (Thm.cterm_of ctxt a)
67370521c09c updated for Isabelle 2015
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3229
diff changeset
   861
      |> Thm.forall_intr (Thm.cterm_of ctxt P)
2665
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   862
  in
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   863
    simple_induct_rule
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   864
  end
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   865
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   866
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   867
(* FIXME: broken by design *)
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   868
fun mk_domain_intro ctxt (Globals {domT, ...}) R R_cases clause =
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   869
  let
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   870
    val ClauseInfo {cdata = ClauseContext {gs, lhs, cqs, ...},
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   871
      qglr = (oqs, _, _, _), ...} = clause
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   872
    val goal = HOLogic.mk_Trueprop (mk_acc domT R $ lhs)
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   873
      |> fold_rev (curry Logic.mk_implies) gs
3239
67370521c09c updated for Isabelle 2015
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3229
diff changeset
   874
      |> Thm.cterm_of ctxt
2665
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   875
  in
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   876
    Goal.init goal
3239
67370521c09c updated for Isabelle 2015
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3229
diff changeset
   877
    |> (SINGLE (resolve_tac ctxt [accI] 1)) |> the
67370521c09c updated for Isabelle 2015
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3229
diff changeset
   878
    |> (SINGLE (eresolve_tac ctxt [Thm.forall_elim_vars 0 R_cases] 1))  |> the
2788
036a19936feb updated to new Isabelle
Christian Urban <urbanc@in.tum.de>
parents: 2781
diff changeset
   879
    |> (SINGLE (auto_tac ctxt)) |> the
2665
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   880
    |> Goal.conclude
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   881
    |> fold_rev forall_intr_rename (map fst oqs ~~ cqs)
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   882
  end
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   883
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   884
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   885
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   886
(** Termination rule **)
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   887
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   888
val wf_induct_rule = @{thm Wellfounded.wfP_induct_rule}
3229
b52e8651591f updated to Isabelle changes
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3227
diff changeset
   889
val wf_in_rel = @{thm Fun_Def.wf_in_rel}
b52e8651591f updated to Isabelle changes
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3227
diff changeset
   890
val in_rel_def = @{thm Fun_Def.in_rel_def}
2665
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   891
3239
67370521c09c updated for Isabelle 2015
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3229
diff changeset
   892
fun mk_nest_term_case ctxt globals R' ihyp clause =
2665
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   893
  let
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   894
    val Globals {z, ...} = globals
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   895
    val ClauseInfo {cdata = ClauseContext {qs, cqs, ags, lhs, case_hyp, ...}, tree,
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   896
      qglr=(oqs, _, _, _), ...} = clause
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   897
3218
89158f401b07 updated to simplifier changes
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3200
diff changeset
   898
    val ih_case =
3239
67370521c09c updated for Isabelle 2015
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3229
diff changeset
   899
      full_simplify (put_simpset HOL_basic_ss ctxt addsimps [case_hyp])
3227
35bb5b013f0e updated with current AFP version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3220
diff changeset
   900
        ihyp
2665
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   901
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   902
    fun step (fixes, assumes) (_ $ arg) u (sub,(hyps,thms)) =
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   903
      let
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   904
        val used = (u @ sub)
3239
67370521c09c updated for Isabelle 2015
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3229
diff changeset
   905
          |> map (fn (ctx,thm) => Function_Context_Tree.export_thm ctxt ctx thm)
2665
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   906
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   907
        val hyp = HOLogic.mk_Trueprop (R' $ arg $ lhs)
3239
67370521c09c updated for Isabelle 2015
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3229
diff changeset
   908
          |> fold_rev (curry Logic.mk_implies o Thm.prop_of) used (* additional hyps *)
67370521c09c updated for Isabelle 2015
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3229
diff changeset
   909
          |> Function_Context_Tree.export_term (fixes, assumes)
67370521c09c updated for Isabelle 2015
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3229
diff changeset
   910
          |> fold_rev (curry Logic.mk_implies o Thm.prop_of) ags
2665
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   911
          |> fold_rev mk_forall_rename (map fst oqs ~~ qs)
3239
67370521c09c updated for Isabelle 2015
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3229
diff changeset
   912
          |> Thm.cterm_of ctxt
2665
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   913
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   914
        val thm = Thm.assume hyp
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   915
          |> fold Thm.forall_elim cqs
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   916
          |> fold Thm.elim_implies ags
3239
67370521c09c updated for Isabelle 2015
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3229
diff changeset
   917
          |> Function_Context_Tree.import_thm ctxt (fixes, assumes)
2665
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   918
          |> fold Thm.elim_implies used (*  "(arg, lhs) : R'"  *)
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   919
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   920
        val z_eq_arg = HOLogic.mk_Trueprop (mk_eq (z, arg))
3239
67370521c09c updated for Isabelle 2015
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3229
diff changeset
   921
          |> Thm.cterm_of ctxt |> Thm.assume
2665
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   922
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   923
        val acc = thm COMP ih_case
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   924
        val z_acc_local = acc
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   925
          |> Conv.fconv_rule
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   926
              (Conv.arg_conv (Conv.arg_conv (K (Thm.symmetric (z_eq_arg RS eq_reflection)))))
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   927
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   928
        val ethm = z_acc_local
3239
67370521c09c updated for Isabelle 2015
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3229
diff changeset
   929
          |> Function_Context_Tree.export_thm ctxt (fixes,
2665
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   930
               z_eq_arg :: case_hyp :: ags @ assumes)
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   931
          |> fold_rev forall_intr_rename (map fst oqs ~~ cqs)
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   932
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   933
        val sub' = sub @ [(([],[]), acc)]
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   934
      in
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   935
        (sub', (hyp :: hyps, ethm :: thms))
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   936
      end
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   937
      | step _ _ _ _ = raise Match
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   938
  in
3239
67370521c09c updated for Isabelle 2015
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3229
diff changeset
   939
    Function_Context_Tree.traverse_tree step tree
2665
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   940
  end
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   941
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   942
3239
67370521c09c updated for Isabelle 2015
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3229
diff changeset
   943
fun mk_nest_term_rule ctxt globals R R_cases clauses =
2665
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   944
  let
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   945
    val Globals { domT, x, z, ... } = globals
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   946
    val acc_R = mk_acc domT R
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   947
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   948
    val R' = Free ("R", fastype_of R)
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   949
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   950
    val Rrel = Free ("R", HOLogic.mk_setT (HOLogic.mk_prodT (domT, domT)))
3229
b52e8651591f updated to Isabelle changes
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3227
diff changeset
   951
    val inrel_R = Const (@{const_name Fun_Def.in_rel},
2665
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   952
      HOLogic.mk_setT (HOLogic.mk_prodT (domT, domT)) --> fastype_of R) $ Rrel
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   953
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   954
    val wfR' = HOLogic.mk_Trueprop (Const (@{const_name Wellfounded.wfP},
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   955
      (domT --> domT --> boolT) --> boolT) $ R')
3239
67370521c09c updated for Isabelle 2015
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3229
diff changeset
   956
      |> Thm.cterm_of ctxt (* "wf R'" *)
2665
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   957
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   958
    (* Inductive Hypothesis: !!z. (z,x):R' ==> z : acc R *)
3108
61db5ad429bb updated to Isabelle 16 January
Christian Urban <urbanc@in.tum.de>
parents: 3045
diff changeset
   959
    val ihyp = Logic.all_const domT $ Abs ("z", domT,
2665
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   960
      Logic.mk_implies (HOLogic.mk_Trueprop (R' $ Bound 0 $ x),
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   961
        HOLogic.mk_Trueprop (acc_R $ Bound 0)))
3239
67370521c09c updated for Isabelle 2015
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3229
diff changeset
   962
      |> Thm.cterm_of ctxt
2665
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   963
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   964
    val ihyp_a = Thm.assume ihyp |> Thm.forall_elim_vars 0
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   965
3239
67370521c09c updated for Isabelle 2015
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3229
diff changeset
   966
    val R_z_x = Thm.cterm_of ctxt (HOLogic.mk_Trueprop (R $ z $ x))
2665
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   967
3239
67370521c09c updated for Isabelle 2015
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3229
diff changeset
   968
    val (hyps, cases) = fold (mk_nest_term_case ctxt globals R' ihyp_a) clauses ([], [])
2665
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   969
  in
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   970
    R_cases
3239
67370521c09c updated for Isabelle 2015
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3229
diff changeset
   971
    |> Thm.forall_elim (Thm.cterm_of ctxt z)
67370521c09c updated for Isabelle 2015
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3229
diff changeset
   972
    |> Thm.forall_elim (Thm.cterm_of ctxt x)
67370521c09c updated for Isabelle 2015
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3229
diff changeset
   973
    |> Thm.forall_elim (Thm.cterm_of ctxt (acc_R $ z))
2665
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   974
    |> curry op COMP (Thm.assume R_z_x)
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   975
    |> fold_rev (curry op COMP) cases
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   976
    |> Thm.implies_intr R_z_x
3239
67370521c09c updated for Isabelle 2015
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3229
diff changeset
   977
    |> Thm.forall_intr (Thm.cterm_of ctxt z)
2665
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   978
    |> (fn it => it COMP accI)
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   979
    |> Thm.implies_intr ihyp
3239
67370521c09c updated for Isabelle 2015
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3229
diff changeset
   980
    |> Thm.forall_intr (Thm.cterm_of ctxt x)
3220
87dbeba4b25a updated to ne Isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3219
diff changeset
   981
    |> (fn it => Drule.compose (it, 2, wf_induct_rule))
2665
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   982
    |> curry op RS (Thm.assume wfR')
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   983
    |> forall_intr_vars
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   984
    |> (fn it => it COMP allI)
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   985
    |> fold Thm.implies_intr hyps
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   986
    |> Thm.implies_intr wfR'
3239
67370521c09c updated for Isabelle 2015
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3229
diff changeset
   987
    |> Thm.forall_intr (Thm.cterm_of ctxt R')
67370521c09c updated for Isabelle 2015
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3229
diff changeset
   988
    |> Thm.forall_elim (Thm.cterm_of ctxt (inrel_R))
2665
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   989
    |> curry op RS wf_in_rel
3239
67370521c09c updated for Isabelle 2015
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3229
diff changeset
   990
    |> full_simplify (put_simpset HOL_basic_ss ctxt addsimps [in_rel_def])
67370521c09c updated for Isabelle 2015
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3229
diff changeset
   991
    |> Thm.forall_intr (Thm.cterm_of ctxt Rrel)
2665
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   992
  end
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   993
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   994
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   995
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   996
(* nominal *)
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   997
fun prepare_nominal_function config defname [((fname, fT), mixfix)] abstract_qglrs lthy =
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   998
  let
2821
c7d4bd9e89e0 fixed problem with earlier commit about nominal_function_common; added facility for specifying an invariant - added a definition of frees_set which need a finiteness invariant
Christian Urban <urbanc@in.tum.de>
parents: 2819
diff changeset
   999
    val NominalFunctionConfig {domintros, default=default_opt, inv=invariant_opt,...} = config
2665
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1000
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1001
    val default_str = the_default "%x. undefined" default_opt (*FIXME dynamic scoping*)
2821
c7d4bd9e89e0 fixed problem with earlier commit about nominal_function_common; added facility for specifying an invariant - added a definition of frees_set which need a finiteness invariant
Christian Urban <urbanc@in.tum.de>
parents: 2819
diff changeset
  1002
    val invariant_str = the_default "%x y. True" invariant_opt
2665
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1003
    val fvar = Free (fname, fT)
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1004
    val domT = domain_type fT
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1005
    val ranT = range_type fT
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1006
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1007
    val default = Syntax.parse_term lthy default_str
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1008
      |> Type.constraint fT |> Syntax.check_term lthy
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1009
2821
c7d4bd9e89e0 fixed problem with earlier commit about nominal_function_common; added facility for specifying an invariant - added a definition of frees_set which need a finiteness invariant
Christian Urban <urbanc@in.tum.de>
parents: 2819
diff changeset
  1010
    val invariant_trm = Syntax.parse_term lthy invariant_str
c7d4bd9e89e0 fixed problem with earlier commit about nominal_function_common; added facility for specifying an invariant - added a definition of frees_set which need a finiteness invariant
Christian Urban <urbanc@in.tum.de>
parents: 2819
diff changeset
  1011
      |> Type.constraint ([domT, ranT] ---> @{typ bool}) |> Syntax.check_term lthy
c7d4bd9e89e0 fixed problem with earlier commit about nominal_function_common; added facility for specifying an invariant - added a definition of frees_set which need a finiteness invariant
Christian Urban <urbanc@in.tum.de>
parents: 2819
diff changeset
  1012
    
2665
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1013
    val (globals, ctxt') = fix_globals domT ranT fvar lthy
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1014
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1015
    val Globals { x, h, ... } = globals
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1016
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1017
    val clauses = map (mk_clause_context x ctxt') abstract_qglrs
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1018
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1019
    val n = length abstract_qglrs
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1020
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1021
    fun build_tree (ClauseContext { ctxt, rhs, ...}) =
3239
67370521c09c updated for Isabelle 2015
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3229
diff changeset
  1022
       Function_Context_Tree.mk_tree (fname, fT) h ctxt rhs
2665
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1023
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1024
    val trees = map build_tree clauses
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1025
    val RCss = map find_calls trees
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1026
2707
747ebf2f066d made eqvt-proof explicit in the function definitions
Christian Urban <urbanc@in.tum.de>
parents: 2677
diff changeset
  1027
    val ((G, GIntro_thms, G_elim, G_induct), lthy) =
3245
017e33849f4d updated to Isabelle 2016-1
Christian Urban <urbanc@in.tum.de>
parents: 3244
diff changeset
  1028
      PROFILE "def_graph" (define_graph (defname ^ "_graph") fvar domT ranT clauses RCss) lthy
2665
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1029
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1030
    val ((f, (_, f_defthm)), lthy) =
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1031
      PROFILE "def_fun" (define_function (defname ^ "_sumC_def") (fname, mixfix) domT ranT G default) lthy
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1032
3239
67370521c09c updated for Isabelle 2015
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3229
diff changeset
  1033
    val RCss = map (map (inst_RC lthy fvar f)) RCss
67370521c09c updated for Isabelle 2015
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3229
diff changeset
  1034
    val trees = map (Function_Context_Tree.inst_tree lthy fvar f) trees
2665
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1035
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1036
    val ((R, RIntro_thmss, R_elim), lthy) =
3245
017e33849f4d updated to Isabelle 2016-1
Christian Urban <urbanc@in.tum.de>
parents: 3244
diff changeset
  1037
      PROFILE "def_rel" (define_recursion_relation (defname ^ "_rel") domT abstract_qglrs clauses RCss) lthy
2665
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1038
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1039
    val (_, lthy) =
3245
017e33849f4d updated to Isabelle 2016-1
Christian Urban <urbanc@in.tum.de>
parents: 3244
diff changeset
  1040
      Local_Theory.abbrev Syntax.mode_default ((Binding.name (defname ^ "_dom"), NoSyn), mk_acc domT R) lthy
2665
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1041
3045
d0ad264f8c4f updated to Isabelle 3 Nov; it includes a hack to work around a bug in the localised version of the quotient package
Christian Urban <urbanc@in.tum.de>
parents: 2994
diff changeset
  1042
    val newthy = Proof_Context.theory_of lthy
2665
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1043
    val clauses = map (transfer_clause_ctx newthy) clauses
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1044
3239
67370521c09c updated for Isabelle 2015
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3229
diff changeset
  1045
    val cert = Thm.cterm_of lthy
2665
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1046
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1047
    val xclauses = PROFILE "xclauses"
3239
67370521c09c updated for Isabelle 2015
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3229
diff changeset
  1048
      (@{map 7} (mk_clause_info globals G f) (1 upto n) clauses abstract_qglrs trees
2665
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1049
        RCss GIntro_thms) RIntro_thmss
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1050
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1051
    val complete =
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1052
      mk_completeness globals clauses abstract_qglrs |> cert |> Thm.assume
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1053
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1054
    val compat =
2821
c7d4bd9e89e0 fixed problem with earlier commit about nominal_function_common; added facility for specifying an invariant - added a definition of frees_set which need a finiteness invariant
Christian Urban <urbanc@in.tum.de>
parents: 2819
diff changeset
  1055
      mk_compat_proof_obligations domT ranT fvar f RCss invariant_trm abstract_qglrs 
2665
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1056
      |> map (cert #> Thm.assume)
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1057
2707
747ebf2f066d made eqvt-proof explicit in the function definitions
Christian Urban <urbanc@in.tum.de>
parents: 2677
diff changeset
  1058
    val G_eqvt = mk_eqvt G |> cert |> Thm.assume
747ebf2f066d made eqvt-proof explicit in the function definitions
Christian Urban <urbanc@in.tum.de>
parents: 2677
diff changeset
  1059
 
2821
c7d4bd9e89e0 fixed problem with earlier commit about nominal_function_common; added facility for specifying an invariant - added a definition of frees_set which need a finiteness invariant
Christian Urban <urbanc@in.tum.de>
parents: 2819
diff changeset
  1060
    val invariant = mk_invariant globals G invariant_trm |> cert |> Thm.assume
c7d4bd9e89e0 fixed problem with earlier commit about nominal_function_common; added facility for specifying an invariant - added a definition of frees_set which need a finiteness invariant
Christian Urban <urbanc@in.tum.de>
parents: 2819
diff changeset
  1061
 
2665
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1062
    val compat_store = store_compat_thms n compat
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1063
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1064
    val (goalstate, values) = PROFILE "prove_stuff"
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1065
      (prove_stuff lthy globals G f R xclauses complete compat
2821
c7d4bd9e89e0 fixed problem with earlier commit about nominal_function_common; added facility for specifying an invariant - added a definition of frees_set which need a finiteness invariant
Christian Urban <urbanc@in.tum.de>
parents: 2819
diff changeset
  1066
         compat_store G_elim G_eqvt invariant) f_defthm
2665
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1067
     
3244
a44479bde681 fixed a problem with two example theories
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3239
diff changeset
  1068
    fun mk_partial_rules newctxt provedgoal =
2665
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1069
      let
2974
b95a2065aa10 generated the partial eqvt-theorem for functions
Christian Urban <urbanc@in.tum.de>
parents: 2973
diff changeset
  1070
        val ((graph_is_function, complete_thm), graph_is_eqvt) =
2665
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1071
          provedgoal
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1072
          |> Conjunction.elim
2974
b95a2065aa10 generated the partial eqvt-theorem for functions
Christian Urban <urbanc@in.tum.de>
parents: 2973
diff changeset
  1073
          |>> fst o Conjunction.elim
b95a2065aa10 generated the partial eqvt-theorem for functions
Christian Urban <urbanc@in.tum.de>
parents: 2973
diff changeset
  1074
          |>> Conjunction.elim
b95a2065aa10 generated the partial eqvt-theorem for functions
Christian Urban <urbanc@in.tum.de>
parents: 2973
diff changeset
  1075
          |>> apfst (Thm.forall_elim_vars 0)
2665
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1076
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1077
        val f_iff = graph_is_function RS (f_defthm RS ex1_implies_iff)
2974
b95a2065aa10 generated the partial eqvt-theorem for functions
Christian Urban <urbanc@in.tum.de>
parents: 2973
diff changeset
  1078
        val f_eqvt = graph_is_function RS (graph_is_eqvt RS (f_defthm RS @{thm fundef_ex1_eqvt}))
2973
d1038e67923a added a flag (eqvt) to termination proofs arising fron nominal_primrecs
Christian Urban <urbanc@in.tum.de>
parents: 2885
diff changeset
  1079
2665
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1080
        val psimps = PROFILE "Proving simplification rules"
3239
67370521c09c updated for Isabelle 2015
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3229
diff changeset
  1081
          (mk_psimps newctxt globals R xclauses values f_iff) graph_is_function
2665
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1082
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1083
        val simple_pinduct = PROFILE "Proving partial induction rule"
3239
67370521c09c updated for Isabelle 2015
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3229
diff changeset
  1084
          (mk_partial_induct_rule newctxt globals R complete_thm) xclauses
2665
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1085
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1086
        val total_intro = PROFILE "Proving nested termination rule"
3239
67370521c09c updated for Isabelle 2015
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3229
diff changeset
  1087
          (mk_nest_term_rule newctxt globals R R_elim) xclauses
2665
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1088
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1089
        val dom_intros =
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1090
          if domintros then SOME (PROFILE "Proving domain introduction rules"
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1091
             (map (mk_domain_intro lthy globals R R_elim)) xclauses)
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1092
           else NONE
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1093
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1094
      in
2973
d1038e67923a added a flag (eqvt) to termination proofs arising fron nominal_primrecs
Christian Urban <urbanc@in.tum.de>
parents: 2885
diff changeset
  1095
        NominalFunctionResult {fs=[f], G=G, R=R, cases=complete_thm,
2665
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1096
          psimps=psimps, simple_pinducts=[simple_pinduct],
2973
d1038e67923a added a flag (eqvt) to termination proofs arising fron nominal_primrecs
Christian Urban <urbanc@in.tum.de>
parents: 2885
diff changeset
  1097
          termination=total_intro, domintros=dom_intros, eqvts=[f_eqvt]}
2665
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1098
      end
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1099
  in
3197
25d11b449e92 definition of an auxiliary graph in nominal-primrec definitions
Christian Urban <urbanc@in.tum.de>
parents: 3108
diff changeset
  1100
    ((f, G, GIntro_thms, G_induct, goalstate, mk_partial_rules), lthy)
2665
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1101
  end
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1102
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1103
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1104
end