src/HOL/Tools/Function/function_core.ML
author wenzelm
Tue, 08 Sep 2015 13:52:13 +0200
changeset 61133 c6d62d092246
parent 61121 efe8b18306b7
child 61340 ce74c00de6b7
permissions -rw-r--r--
clarified JEDIT_JAVA_OPTIONS: separate defaults for 32 vs. 64 bit;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
33099
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
     1
(*  Title:      HOL/Tools/Function/function_core.ML
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
     2
    Author:     Alexander Krauss, TU Muenchen
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
     3
41114
f9ae7c2abf7e tuned headers
krauss
parents: 39288
diff changeset
     4
Core of the function package.
33099
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
     5
*)
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
     6
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
     7
signature FUNCTION_CORE =
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
     8
sig
34232
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
     9
  val trace: bool Unsynchronized.ref
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
    10
  val prepare_function : Function_Common.function_config
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
    11
    -> string (* defname *)
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
    12
    -> ((bstring * typ) * mixfix) list (* defined symbol *)
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
    13
    -> ((bstring * typ) list * term list * term * term) list (* specification *)
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
    14
    -> local_theory
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
    15
    -> (term   (* f *)
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
    16
        * thm  (* goalstate *)
60643
9173467ec5b6 clarified context;
wenzelm
parents: 59880
diff changeset
    17
        * (Proof.context -> thm -> Function_Common.function_result) (* continuation *)
34232
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
    18
       ) * local_theory
33099
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
    19
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
    20
end
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
    21
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
    22
structure Function_Core : FUNCTION_CORE =
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
    23
struct
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
    24
34232
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
    25
val trace = Unsynchronized.ref false
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
    26
fun trace_msg msg = if ! trace then tracing (msg ()) else ()
33099
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
    27
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
    28
val boolT = HOLogic.boolT
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
    29
val mk_eq = HOLogic.mk_eq
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
    30
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
    31
open Function_Lib
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
    32
open Function_Common
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
    33
34232
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
    34
datatype globals = Globals of
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
    35
 {fvar: term,
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
    36
  domT: typ,
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
    37
  ranT: typ,
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
    38
  h: term,
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
    39
  y: term,
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
    40
  x: term,
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
    41
  z: term,
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
    42
  a: term,
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
    43
  P: term,
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
    44
  D: term,
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
    45
  Pbool:term}
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
    46
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
    47
datatype rec_call_info = RCInfo of
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
    48
 {RIvs: (string * typ) list,  (* Call context: fixes and assumes *)
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
    49
  CCas: thm list,
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
    50
  rcarg: term,                 (* The recursive argument *)
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
    51
  llRI: thm,
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
    52
  h_assum: term}
33099
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
    53
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
    54
34232
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
    55
datatype clause_context = ClauseContext of
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
    56
 {ctxt : Proof.context,
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
    57
  qs : term list,
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
    58
  gs : term list,
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
    59
  lhs: term,
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
    60
  rhs: term,
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
    61
  cqs: cterm list,
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
    62
  ags: thm list,
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
    63
  case_hyp : thm}
33099
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
    64
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
    65
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
    66
fun transfer_clause_ctx thy (ClauseContext { ctxt, qs, gs, lhs, rhs, cqs, ags, case_hyp }) =
42361
23f352990944 modernized structure Proof_Context;
wenzelm
parents: 41846
diff changeset
    67
  ClauseContext { ctxt = Proof_Context.transfer thy ctxt,
34232
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
    68
    qs = qs, gs = gs, lhs = lhs, rhs = rhs, cqs = cqs, ags = ags, case_hyp = case_hyp }
33099
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
    69
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
    70
34232
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
    71
datatype clause_info = ClauseInfo of
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
    72
 {no: int,
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
    73
  qglr : ((string * typ) list * term list * term * term),
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
    74
  cdata : clause_context,
58816
aab139c0003f modernized setup;
wenzelm
parents: 58634
diff changeset
    75
  tree: Function_Context_Tree.ctx_tree,
34232
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
    76
  lGI: thm,
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
    77
  RCs: rec_call_info list}
33099
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
    78
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
    79
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
    80
(* Theory dependencies. *)
34232
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
    81
val acc_induct_rule = @{thm accp_induct_rule}
33099
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
    82
55085
0e8e4dc55866 moved 'fundef_cong' attribute (and other basic 'fun' stuff) up the dependency chain
blanchet
parents: 53603
diff changeset
    83
val ex1_implies_ex = @{thm Fun_Def.fundef_ex1_existence}
0e8e4dc55866 moved 'fundef_cong' attribute (and other basic 'fun' stuff) up the dependency chain
blanchet
parents: 53603
diff changeset
    84
val ex1_implies_un = @{thm Fun_Def.fundef_ex1_uniqueness}
0e8e4dc55866 moved 'fundef_cong' attribute (and other basic 'fun' stuff) up the dependency chain
blanchet
parents: 53603
diff changeset
    85
val ex1_implies_iff = @{thm Fun_Def.fundef_ex1_iff}
33099
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
    86
34232
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
    87
val acc_downward = @{thm accp_downward}
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
    88
val accI = @{thm accp.accI}
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
    89
val case_split = @{thm HOL.case_split}
55085
0e8e4dc55866 moved 'fundef_cong' attribute (and other basic 'fun' stuff) up the dependency chain
blanchet
parents: 53603
diff changeset
    90
val fundef_default_value = @{thm Fun_Def.fundef_default_value}
34232
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
    91
val not_acc_down = @{thm not_accp_down}
33099
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
    92
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
    93
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
    94
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
    95
fun find_calls tree =
34232
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
    96
  let
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
    97
    fun add_Ri (fixes,assumes) (_ $ arg) _ (_, xs) =
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
    98
      ([], (fixes, assumes, arg) :: xs)
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
    99
      | add_Ri _ _ _ _ = raise Match
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
   100
  in
58816
aab139c0003f modernized setup;
wenzelm
parents: 58634
diff changeset
   101
    rev (Function_Context_Tree.traverse_tree add_Ri tree [])
34232
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
   102
  end
33099
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   103
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   104
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   105
(** building proof obligations *)
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   106
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   107
fun mk_compat_proof_obligations domT ranT fvar f glrs =
34232
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
   108
  let
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
   109
    fun mk_impl ((qs, gs, lhs, rhs),(qs', gs', lhs', rhs')) =
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
   110
      let
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
   111
        val shift = incr_boundvars (length qs')
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
   112
      in
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
   113
        Logic.mk_implies
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
   114
          (HOLogic.mk_Trueprop (HOLogic.eq_const domT $ shift lhs $ lhs'),
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
   115
            HOLogic.mk_Trueprop (HOLogic.eq_const ranT $ shift rhs $ rhs'))
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
   116
        |> fold_rev (curry Logic.mk_implies) (map shift gs @ gs')
46217
7b19666f0e3d renamed Term.all to Logic.all_const, in accordance to HOLogic.all_const;
wenzelm
parents: 42793
diff changeset
   117
        |> fold_rev (fn (n,T) => fn b => Logic.all_const T $ Abs(n,T,b)) (qs @ qs')
34232
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
   118
        |> curry abstract_over fvar
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
   119
        |> curry subst_bound f
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
   120
      end
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
   121
  in
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
   122
    map mk_impl (unordered_pairs glrs)
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
   123
  end
33099
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   124
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   125
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   126
fun mk_completeness (Globals {x, Pbool, ...}) clauses qglrs =
34232
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
   127
  let
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
   128
    fun mk_case (ClauseContext {qs, gs, lhs, ...}, (oqs, _, _, _)) =
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
   129
      HOLogic.mk_Trueprop Pbool
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
   130
      |> curry Logic.mk_implies (HOLogic.mk_Trueprop (mk_eq (x, lhs)))
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
   131
      |> fold_rev (curry Logic.mk_implies) gs
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
   132
      |> fold_rev mk_forall_rename (map fst oqs ~~ qs)
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
   133
  in
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
   134
    HOLogic.mk_Trueprop Pbool
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
   135
    |> fold_rev (curry Logic.mk_implies o mk_case) (clauses ~~ qglrs)
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
   136
    |> mk_forall_rename ("x", x)
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
   137
    |> mk_forall_rename ("P", Pbool)
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
   138
  end
33099
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   139
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   140
(** making a context with it's own local bindings **)
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   141
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   142
fun mk_clause_context x ctxt (pre_qs,pre_gs,pre_lhs,pre_rhs) =
34232
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
   143
  let
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
   144
    val (qs, ctxt') = Variable.variant_fixes (map fst pre_qs) ctxt
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
   145
      |>> map2 (fn (_, T) => fn n => Free (n, T)) pre_qs
33099
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   146
34232
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
   147
    fun inst t = subst_bounds (rev qs, t)
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
   148
    val gs = map inst pre_gs
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
   149
    val lhs = inst pre_lhs
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
   150
    val rhs = inst pre_rhs
33099
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   151
59621
291934bac95e Thm.cterm_of and Thm.ctyp_of operate on local context;
wenzelm
parents: 59618
diff changeset
   152
    val cqs = map (Thm.cterm_of ctxt') qs
291934bac95e Thm.cterm_of and Thm.ctyp_of operate on local context;
wenzelm
parents: 59618
diff changeset
   153
    val ags = map (Thm.assume o Thm.cterm_of ctxt') gs
33099
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   154
59618
e6939796717e clarified context;
wenzelm
parents: 59582
diff changeset
   155
    val case_hyp =
59621
291934bac95e Thm.cterm_of and Thm.ctyp_of operate on local context;
wenzelm
parents: 59618
diff changeset
   156
      Thm.assume (Thm.cterm_of ctxt' (HOLogic.mk_Trueprop (mk_eq (x, lhs))))
34232
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
   157
  in
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
   158
    ClauseContext { ctxt = ctxt', qs = qs, gs = gs, lhs = lhs, rhs = rhs,
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
   159
      cqs = cqs, ags = ags, case_hyp = case_hyp }
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
   160
  end
33099
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   161
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   162
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   163
(* lowlevel term function. FIXME: remove *)
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   164
fun abstract_over_list vs body =
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   165
  let
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   166
    fun abs lev v tm =
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   167
      if v aconv tm then Bound lev
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   168
      else
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   169
        (case tm of
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   170
          Abs (a, T, t) => Abs (a, T, abs (lev + 1) v t)
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   171
        | t $ u => abs lev v t $ abs lev v u
34232
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
   172
        | t => t)
33099
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   173
  in
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   174
    fold_index (fn (i, v) => fn t => abs i v t) vs body
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   175
  end
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   176
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   177
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   178
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   179
fun mk_clause_info globals G f no cdata qglr tree RCs GIntro_thm RIntro_thms =
34232
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
   180
  let
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
   181
    val Globals {h, ...} = globals
33099
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   182
34232
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
   183
    val ClauseContext { ctxt, qs, cqs, ags, ... } = cdata
33099
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   184
34232
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
   185
    (* Instantiate the GIntro thm with "f" and import into the clause context. *)
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
   186
    val lGI = GIntro_thm
59621
291934bac95e Thm.cterm_of and Thm.ctyp_of operate on local context;
wenzelm
parents: 59618
diff changeset
   187
      |> Thm.forall_elim (Thm.cterm_of ctxt f)
36945
9bec62c10714 less pervasive names from structure Thm;
wenzelm
parents: 36936
diff changeset
   188
      |> fold Thm.forall_elim cqs
34232
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
   189
      |> fold Thm.elim_implies ags
33099
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   190
34232
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
   191
    fun mk_call_info (rcfix, rcassm, rcarg) RI =
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
   192
      let
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
   193
        val llRI = RI
36945
9bec62c10714 less pervasive names from structure Thm;
wenzelm
parents: 36936
diff changeset
   194
          |> fold Thm.forall_elim cqs
59621
291934bac95e Thm.cterm_of and Thm.ctyp_of operate on local context;
wenzelm
parents: 59618
diff changeset
   195
          |> fold (Thm.forall_elim o Thm.cterm_of ctxt o Free) rcfix
34232
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
   196
          |> fold Thm.elim_implies ags
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
   197
          |> fold Thm.elim_implies rcassm
33099
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   198
34232
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
   199
        val h_assum =
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
   200
          HOLogic.mk_Trueprop (G $ rcarg $ (h $ rcarg))
59582
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59580
diff changeset
   201
          |> fold_rev (curry Logic.mk_implies o Thm.prop_of) rcassm
34232
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
   202
          |> fold_rev (Logic.all o Free) rcfix
42361
23f352990944 modernized structure Proof_Context;
wenzelm
parents: 41846
diff changeset
   203
          |> Pattern.rewrite_term (Proof_Context.theory_of ctxt) [(f, h)] []
34232
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
   204
          |> abstract_over_list (rev qs)
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
   205
      in
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
   206
        RCInfo {RIvs=rcfix, rcarg=rcarg, CCas=rcassm, llRI=llRI, h_assum=h_assum}
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
   207
      end
33099
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   208
34232
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
   209
    val RC_infos = map2 mk_call_info RCs RIntro_thms
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
   210
  in
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
   211
    ClauseInfo {no=no, cdata=cdata, qglr=qglr, lGI=lGI, RCs=RC_infos,
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
   212
      tree=tree}
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
   213
  end
33099
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   214
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   215
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   216
fun store_compat_thms 0 thms = []
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   217
  | store_compat_thms n thms =
34232
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
   218
  let
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
   219
    val (thms1, thms2) = chop n thms
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
   220
  in
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
   221
    (thms1 :: store_compat_thms (n - 1) thms2)
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
   222
  end
33099
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   223
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   224
(* expects i <= j *)
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   225
fun lookup_compat_thm i j cts =
34232
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
   226
  nth (nth cts (i - 1)) (j - i)
33099
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   227
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   228
(* Returns "Gsi, Gsj, lhs_i = lhs_j |-- rhs_j_f = rhs_i_f" *)
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   229
(* if j < i, then turn around *)
59627
bb1e4a35d506 clarified context;
wenzelm
parents: 59621
diff changeset
   230
fun get_compat_thm ctxt cts i j ctxi ctxj =
34232
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
   231
  let
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
   232
    val ClauseContext {cqs=cqsi,ags=agsi,lhs=lhsi,...} = ctxi
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
   233
    val ClauseContext {cqs=cqsj,ags=agsj,lhs=lhsj,...} = ctxj
33099
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   234
59627
bb1e4a35d506 clarified context;
wenzelm
parents: 59621
diff changeset
   235
    val lhsi_eq_lhsj = Thm.cterm_of ctxt (HOLogic.mk_Trueprop (mk_eq (lhsi, lhsj)))
34232
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
   236
  in if j < i then
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
   237
    let
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
   238
      val compat = lookup_compat_thm j i cts
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
   239
    in
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
   240
      compat         (* "!!qj qi. Gsj => Gsi => lhsj = lhsi ==> rhsj = rhsi" *)
36945
9bec62c10714 less pervasive names from structure Thm;
wenzelm
parents: 36936
diff changeset
   241
      |> fold Thm.forall_elim (cqsj @ cqsi) (* "Gsj => Gsi => lhsj = lhsi ==> rhsj = rhsi" *)
34232
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
   242
      |> fold Thm.elim_implies agsj
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
   243
      |> fold Thm.elim_implies agsi
36945
9bec62c10714 less pervasive names from structure Thm;
wenzelm
parents: 36936
diff changeset
   244
      |> Thm.elim_implies ((Thm.assume lhsi_eq_lhsj) RS sym) (* "Gsj, Gsi, lhsi = lhsj |-- rhsj = rhsi" *)
33099
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   245
    end
34232
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
   246
    else
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
   247
    let
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
   248
      val compat = lookup_compat_thm i j cts
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
   249
    in
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
   250
      compat        (* "!!qi qj. Gsi => Gsj => lhsi = lhsj ==> rhsi = rhsj" *)
36945
9bec62c10714 less pervasive names from structure Thm;
wenzelm
parents: 36936
diff changeset
   251
      |> fold Thm.forall_elim (cqsi @ cqsj) (* "Gsi => Gsj => lhsi = lhsj ==> rhsi = rhsj" *)
34232
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
   252
      |> fold Thm.elim_implies agsi
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
   253
      |> fold Thm.elim_implies agsj
36945
9bec62c10714 less pervasive names from structure Thm;
wenzelm
parents: 36936
diff changeset
   254
      |> Thm.elim_implies (Thm.assume lhsi_eq_lhsj)
34232
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
   255
      |> (fn thm => thm RS sym) (* "Gsi, Gsj, lhsi = lhsj |-- rhsj = rhsi" *)
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
   256
    end
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
   257
  end
33099
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   258
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   259
(* Generates the replacement lemma in fully quantified form. *)
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 49170
diff changeset
   260
fun mk_replacement_lemma ctxt h ih_elim clause =
34232
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
   261
  let
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
   262
    val ClauseInfo {cdata=ClauseContext {qs, lhs, cqs, ags, case_hyp, ...},
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
   263
      RCs, tree, ...} = clause
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
   264
    local open Conv in
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
   265
      val ih_conv = arg1_conv o arg_conv o arg_conv
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
   266
    end
33099
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   267
34232
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
   268
    val ih_elim_case =
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
   269
      Conv.fconv_rule (ih_conv (K (case_hyp RS eq_reflection))) ih_elim
33099
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   270
34232
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
   271
    val Ris = map (fn RCInfo {llRI, ...} => llRI) RCs
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
   272
    val h_assums = map (fn RCInfo {h_assum, ...} =>
59621
291934bac95e Thm.cterm_of and Thm.ctyp_of operate on local context;
wenzelm
parents: 59618
diff changeset
   273
      Thm.assume (Thm.cterm_of ctxt (subst_bounds (rev qs, h_assum)))) RCs
34232
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
   274
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
   275
    val (eql, _) =
58816
aab139c0003f modernized setup;
wenzelm
parents: 58634
diff changeset
   276
      Function_Context_Tree.rewrite_by_tree ctxt h ih_elim_case (Ris ~~ h_assums) tree
33099
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   277
34232
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
   278
    val replace_lemma = (eql RS meta_eq_to_obj_eq)
59582
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59580
diff changeset
   279
      |> Thm.implies_intr (Thm.cprop_of case_hyp)
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59580
diff changeset
   280
      |> fold_rev (Thm.implies_intr o Thm.cprop_of) h_assums
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59580
diff changeset
   281
      |> fold_rev (Thm.implies_intr o Thm.cprop_of) ags
36945
9bec62c10714 less pervasive names from structure Thm;
wenzelm
parents: 36936
diff changeset
   282
      |> fold_rev Thm.forall_intr cqs
34232
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
   283
      |> Thm.close_derivation
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
   284
  in
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
   285
    replace_lemma
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
   286
  end
33099
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   287
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   288
59627
bb1e4a35d506 clarified context;
wenzelm
parents: 59621
diff changeset
   289
fun mk_uniqueness_clause ctxt globals compat_store clausei clausej RLj =
34232
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
   290
  let
59627
bb1e4a35d506 clarified context;
wenzelm
parents: 59621
diff changeset
   291
    val thy = Proof_Context.theory_of ctxt
bb1e4a35d506 clarified context;
wenzelm
parents: 59621
diff changeset
   292
34232
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
   293
    val Globals {h, y, x, fvar, ...} = globals
59582
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59580
diff changeset
   294
    val ClauseInfo {no=i, cdata=cctxi as ClauseContext {ctxt=ctxti, lhs=lhsi, case_hyp, ...}, ...} =
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59580
diff changeset
   295
      clausei
34232
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
   296
    val ClauseInfo {no=j, qglr=cdescj, RCs=RCsj, ...} = clausej
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
   297
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
   298
    val cctxj as ClauseContext {ags = agsj', lhs = lhsj', rhs = rhsj', qs = qsj', cqs = cqsj', ...} =
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
   299
      mk_clause_context x ctxti cdescj
33099
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   300
34232
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
   301
    val rhsj'h = Pattern.rewrite_term thy [(fvar,h)] [] rhsj'
59627
bb1e4a35d506 clarified context;
wenzelm
parents: 59621
diff changeset
   302
    val compat = get_compat_thm ctxt compat_store i j cctxi cctxj
59582
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59580
diff changeset
   303
    val Ghsj' =
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59580
diff changeset
   304
      map (fn RCInfo {h_assum, ...} =>
59627
bb1e4a35d506 clarified context;
wenzelm
parents: 59621
diff changeset
   305
        Thm.assume (Thm.cterm_of ctxt (subst_bounds (rev qsj', h_assum)))) RCsj
33099
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   306
34232
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
   307
    val RLj_import = RLj
36945
9bec62c10714 less pervasive names from structure Thm;
wenzelm
parents: 36936
diff changeset
   308
      |> fold Thm.forall_elim cqsj'
34232
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
   309
      |> fold Thm.elim_implies agsj'
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
   310
      |> fold Thm.elim_implies Ghsj'
33099
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   311
59627
bb1e4a35d506 clarified context;
wenzelm
parents: 59621
diff changeset
   312
    val y_eq_rhsj'h = Thm.assume (Thm.cterm_of ctxt (HOLogic.mk_Trueprop (mk_eq (y, rhsj'h))))
bb1e4a35d506 clarified context;
wenzelm
parents: 59621
diff changeset
   313
    val lhsi_eq_lhsj' = Thm.assume (Thm.cterm_of ctxt (HOLogic.mk_Trueprop (mk_eq (lhsi, lhsj'))))
34232
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
   314
       (* lhs_i = lhs_j' |-- lhs_i = lhs_j' *)
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
   315
  in
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
   316
    (trans OF [case_hyp, lhsi_eq_lhsj']) (* lhs_i = lhs_j' |-- x = lhs_j' *)
36945
9bec62c10714 less pervasive names from structure Thm;
wenzelm
parents: 36936
diff changeset
   317
    |> Thm.implies_elim RLj_import
34232
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
   318
      (* Rj1' ... Rjk', lhs_i = lhs_j' |-- rhs_j'_h = rhs_j'_f *)
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
   319
    |> (fn it => trans OF [it, compat])
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
   320
      (* lhs_i = lhs_j', Gj', Rj1' ... Rjk' |-- rhs_j'_h = rhs_i_f *)
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
   321
    |> (fn it => trans OF [y_eq_rhsj'h, it])
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
   322
      (* lhs_i = lhs_j', Gj', Rj1' ... Rjk', y = rhs_j_h' |-- y = rhs_i_f *)
59582
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59580
diff changeset
   323
    |> fold_rev (Thm.implies_intr o Thm.cprop_of) Ghsj'
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59580
diff changeset
   324
    |> fold_rev (Thm.implies_intr o Thm.cprop_of) agsj'
34232
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
   325
      (* lhs_i = lhs_j' , y = rhs_j_h' |-- Gj', Rj1'...Rjk' ==> y = rhs_i_f *)
59582
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59580
diff changeset
   326
    |> Thm.implies_intr (Thm.cprop_of y_eq_rhsj'h)
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59580
diff changeset
   327
    |> Thm.implies_intr (Thm.cprop_of lhsi_eq_lhsj')
59627
bb1e4a35d506 clarified context;
wenzelm
parents: 59621
diff changeset
   328
    |> fold_rev Thm.forall_intr (Thm.cterm_of ctxt h :: cqsj')
34232
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
   329
  end
33099
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   330
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   331
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   332
59627
bb1e4a35d506 clarified context;
wenzelm
parents: 59621
diff changeset
   333
fun mk_uniqueness_case ctxt
bb1e4a35d506 clarified context;
wenzelm
parents: 59621
diff changeset
   334
    globals G f ihyp ih_intro G_cases compat_store clauses rep_lemmas clausei =
34232
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
   335
  let
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 49170
diff changeset
   336
    val thy = Proof_Context.theory_of ctxt
34232
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
   337
    val Globals {x, y, ranT, fvar, ...} = globals
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
   338
    val ClauseInfo {cdata = ClauseContext {lhs, rhs, cqs, ags, case_hyp, ...}, lGI, RCs, ...} = clausei
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
   339
    val rhsC = Pattern.rewrite_term thy [(fvar, f)] [] rhs
33099
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   340
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 49170
diff changeset
   341
    val ih_intro_case = full_simplify (put_simpset HOL_basic_ss ctxt addsimps [case_hyp]) ih_intro
33099
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   342
34232
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
   343
    fun prep_RC (RCInfo {llRI, RIvs, CCas, ...}) = (llRI RS ih_intro_case)
59582
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59580
diff changeset
   344
      |> fold_rev (Thm.implies_intr o Thm.cprop_of) CCas
59621
291934bac95e Thm.cterm_of and Thm.ctyp_of operate on local context;
wenzelm
parents: 59618
diff changeset
   345
      |> fold_rev (Thm.forall_intr o Thm.cterm_of ctxt o Free) RIvs
34232
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
   346
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
   347
    val existence = fold (curry op COMP o prep_RC) RCs lGI
33099
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   348
59621
291934bac95e Thm.cterm_of and Thm.ctyp_of operate on local context;
wenzelm
parents: 59618
diff changeset
   349
    val P = Thm.cterm_of ctxt (mk_eq (y, rhsC))
291934bac95e Thm.cterm_of and Thm.ctyp_of operate on local context;
wenzelm
parents: 59618
diff changeset
   350
    val G_lhs_y = Thm.assume (Thm.cterm_of ctxt (HOLogic.mk_Trueprop (G $ lhs $ y)))
33099
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   351
34232
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
   352
    val unique_clauses =
59627
bb1e4a35d506 clarified context;
wenzelm
parents: 59621
diff changeset
   353
      map2 (mk_uniqueness_clause ctxt globals compat_store clausei) clauses rep_lemmas
33099
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   354
36270
fd95c0514623 tolerate eta-variants in f_graph.cases (from inductive package); added test case;
krauss
parents: 34232
diff changeset
   355
    fun elim_implies_eta A AB =
58950
d07464875dd4 optional proof context for unify operations, for the sake of proper local options;
wenzelm
parents: 58816
diff changeset
   356
      Thm.bicompose (SOME ctxt) {flatten = false, match = true, incremented = false}
d07464875dd4 optional proof context for unify operations, for the sake of proper local options;
wenzelm
parents: 58816
diff changeset
   357
        (false, A, 0) 1 AB
52223
5bb6ae8acb87 tuned signature -- more explicit flags for low-level Thm.bicompose;
wenzelm
parents: 51717
diff changeset
   358
      |> Seq.list_of |> the_single
36270
fd95c0514623 tolerate eta-variants in f_graph.cases (from inductive package); added test case;
krauss
parents: 34232
diff changeset
   359
34232
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
   360
    val uniqueness = G_cases
59621
291934bac95e Thm.cterm_of and Thm.ctyp_of operate on local context;
wenzelm
parents: 59618
diff changeset
   361
      |> Thm.forall_elim (Thm.cterm_of ctxt lhs)
291934bac95e Thm.cterm_of and Thm.ctyp_of operate on local context;
wenzelm
parents: 59618
diff changeset
   362
      |> Thm.forall_elim (Thm.cterm_of ctxt y)
36945
9bec62c10714 less pervasive names from structure Thm;
wenzelm
parents: 36936
diff changeset
   363
      |> Thm.forall_elim P
34232
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
   364
      |> Thm.elim_implies G_lhs_y
36270
fd95c0514623 tolerate eta-variants in f_graph.cases (from inductive package); added test case;
krauss
parents: 34232
diff changeset
   365
      |> fold elim_implies_eta unique_clauses
59582
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59580
diff changeset
   366
      |> Thm.implies_intr (Thm.cprop_of G_lhs_y)
59621
291934bac95e Thm.cterm_of and Thm.ctyp_of operate on local context;
wenzelm
parents: 59618
diff changeset
   367
      |> Thm.forall_intr (Thm.cterm_of ctxt y)
33099
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   368
59621
291934bac95e Thm.cterm_of and Thm.ctyp_of operate on local context;
wenzelm
parents: 59618
diff changeset
   369
    val P2 = Thm.cterm_of ctxt (lambda y (G $ lhs $ y)) (* P2 y := (lhs, y): G *)
33099
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   370
34232
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
   371
    val exactly_one =
59582
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59580
diff changeset
   372
      @{thm ex1I}
60801
7664e0916eec tuned signature;
wenzelm
parents: 60752
diff changeset
   373
      |> Thm.instantiate'
59621
291934bac95e Thm.cterm_of and Thm.ctyp_of operate on local context;
wenzelm
parents: 59618
diff changeset
   374
          [SOME (Thm.ctyp_of ctxt ranT)]
291934bac95e Thm.cterm_of and Thm.ctyp_of operate on local context;
wenzelm
parents: 59618
diff changeset
   375
          [SOME P2, SOME (Thm.cterm_of ctxt rhsC)]
34232
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
   376
      |> curry (op COMP) existence
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
   377
      |> curry (op COMP) uniqueness
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 49170
diff changeset
   378
      |> simplify (put_simpset HOL_basic_ss ctxt addsimps [case_hyp RS sym])
59582
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59580
diff changeset
   379
      |> Thm.implies_intr (Thm.cprop_of case_hyp)
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59580
diff changeset
   380
      |> fold_rev (Thm.implies_intr o Thm.cprop_of) ags
36945
9bec62c10714 less pervasive names from structure Thm;
wenzelm
parents: 36936
diff changeset
   381
      |> fold_rev Thm.forall_intr cqs
33099
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   382
34232
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
   383
    val function_value =
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
   384
      existence
36945
9bec62c10714 less pervasive names from structure Thm;
wenzelm
parents: 36936
diff changeset
   385
      |> Thm.implies_intr ihyp
59582
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59580
diff changeset
   386
      |> Thm.implies_intr (Thm.cprop_of case_hyp)
59621
291934bac95e Thm.cterm_of and Thm.ctyp_of operate on local context;
wenzelm
parents: 59618
diff changeset
   387
      |> Thm.forall_intr (Thm.cterm_of ctxt x)
291934bac95e Thm.cterm_of and Thm.ctyp_of operate on local context;
wenzelm
parents: 59618
diff changeset
   388
      |> Thm.forall_elim (Thm.cterm_of ctxt lhs)
34232
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
   389
      |> curry (op RS) refl
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
   390
  in
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
   391
    (exactly_one, function_value)
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
   392
  end
33099
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   393
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   394
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   395
fun prove_stuff ctxt globals G f R clauses complete compat compat_store G_elim f_def =
34232
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
   396
  let
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
   397
    val Globals {h, domT, ranT, x, ...} = globals
33099
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   398
34232
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
   399
    (* Inductive Hypothesis: !!z. (z,x):R ==> EX!y. (z,y):G *)
46217
7b19666f0e3d renamed Term.all to Logic.all_const, in accordance to HOLogic.all_const;
wenzelm
parents: 42793
diff changeset
   400
    val ihyp = Logic.all_const domT $ Abs ("z", domT,
34232
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
   401
      Logic.mk_implies (HOLogic.mk_Trueprop (R $ Bound 0 $ x),
38558
32ad17fe2b9c tuned quotes
haftmann
parents: 38549
diff changeset
   402
        HOLogic.mk_Trueprop (Const (@{const_name Ex1}, (ranT --> boolT) --> boolT) $
34232
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
   403
          Abs ("y", ranT, G $ Bound 1 $ Bound 0))))
59621
291934bac95e Thm.cterm_of and Thm.ctyp_of operate on local context;
wenzelm
parents: 59618
diff changeset
   404
      |> Thm.cterm_of ctxt
33099
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   405
36945
9bec62c10714 less pervasive names from structure Thm;
wenzelm
parents: 36936
diff changeset
   406
    val ihyp_thm = Thm.assume ihyp |> Thm.forall_elim_vars 0
34232
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
   407
    val ih_intro = ihyp_thm RS (f_def RS ex1_implies_ex)
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
   408
    val ih_elim = ihyp_thm RS (f_def RS ex1_implies_un)
60801
7664e0916eec tuned signature;
wenzelm
parents: 60752
diff changeset
   409
      |> Thm.instantiate' [] [NONE, SOME (Thm.cterm_of ctxt h)]
33099
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   410
34232
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
   411
    val _ = trace_msg (K "Proving Replacement lemmas...")
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 49170
diff changeset
   412
    val repLemmas = map (mk_replacement_lemma ctxt h ih_elim) clauses
33099
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   413
34232
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
   414
    val _ = trace_msg (K "Proving cases for unique existence...")
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
   415
    val (ex1s, values) =
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 49170
diff changeset
   416
      split_list
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 49170
diff changeset
   417
        (map
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 49170
diff changeset
   418
          (mk_uniqueness_case ctxt globals G f ihyp ih_intro G_elim compat_store clauses repLemmas)
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 49170
diff changeset
   419
          clauses)
33099
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   420
34232
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
   421
    val _ = trace_msg (K "Proving: Graph is a function")
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
   422
    val graph_is_function = complete
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
   423
      |> Thm.forall_elim_vars 0
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
   424
      |> fold (curry op COMP) ex1s
36945
9bec62c10714 less pervasive names from structure Thm;
wenzelm
parents: 36936
diff changeset
   425
      |> Thm.implies_intr (ihyp)
59621
291934bac95e Thm.cterm_of and Thm.ctyp_of operate on local context;
wenzelm
parents: 59618
diff changeset
   426
      |> Thm.implies_intr (Thm.cterm_of ctxt (HOLogic.mk_Trueprop (mk_acc domT R $ x)))
291934bac95e Thm.cterm_of and Thm.ctyp_of operate on local context;
wenzelm
parents: 59618
diff changeset
   427
      |> Thm.forall_intr (Thm.cterm_of ctxt x)
52467
24c6ddb48cb8 tuned signature;
wenzelm
parents: 52456
diff changeset
   428
      |> (fn it => Drule.compose (it, 2, acc_induct_rule)) (* "EX! y. (?x,y):G" *)
59582
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59580
diff changeset
   429
      |> (fn it =>
59621
291934bac95e Thm.cterm_of and Thm.ctyp_of operate on local context;
wenzelm
parents: 59618
diff changeset
   430
          fold (Thm.forall_intr o Thm.cterm_of ctxt o Var)
59618
e6939796717e clarified context;
wenzelm
parents: 59582
diff changeset
   431
            (Term.add_vars (Thm.prop_of it) []) it)
33099
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   432
34232
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
   433
    val goalstate =  Conjunction.intr graph_is_function complete
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
   434
      |> Thm.close_derivation
52456
960202346d0c tuned signature;
wenzelm
parents: 52384
diff changeset
   435
      |> Goal.protect 0
59582
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59580
diff changeset
   436
      |> fold_rev (Thm.implies_intr o Thm.cprop_of) compat
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59580
diff changeset
   437
      |> Thm.implies_intr (Thm.cprop_of complete)
34232
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
   438
  in
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
   439
    (goalstate, values)
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
   440
  end
33099
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   441
33348
bb65583ab70d absorbed inductive_wrap function into Function_Core; more conventional argument order; tuned
krauss
parents: 33278
diff changeset
   442
(* wrapper -- restores quantifiers in rule specifications *)
bb65583ab70d absorbed inductive_wrap function into Function_Core; more conventional argument order; tuned
krauss
parents: 33278
diff changeset
   443
fun inductive_def (binding as ((R, T), _)) intrs lthy =
bb65583ab70d absorbed inductive_wrap function into Function_Core; more conventional argument order; tuned
krauss
parents: 33278
diff changeset
   444
  let
bb65583ab70d absorbed inductive_wrap function into Function_Core; more conventional argument order; tuned
krauss
parents: 33278
diff changeset
   445
    val ({intrs = intrs_gen, elims = [elim_gen], preds = [ Rdef ], induct, ...}, lthy) =
bb65583ab70d absorbed inductive_wrap function into Function_Core; more conventional argument order; tuned
krauss
parents: 33278
diff changeset
   446
      lthy
59880
30687c3f2b10 clarified role of naming for background theory: transform_binding (e.g. for "concealed" flag) uses naming of hypothetical context;
wenzelm
parents: 59859
diff changeset
   447
      |> Proof_Context.concealed
33348
bb65583ab70d absorbed inductive_wrap function into Function_Core; more conventional argument order; tuned
krauss
parents: 33278
diff changeset
   448
      |> Inductive.add_inductive_i
33350
b2b78c5ef771 less verbose inductive invocation
krauss
parents: 33349
diff changeset
   449
          {quiet_mode = true,
b2b78c5ef771 less verbose inductive invocation
krauss
parents: 33349
diff changeset
   450
            verbose = ! trace,
33348
bb65583ab70d absorbed inductive_wrap function into Function_Core; more conventional argument order; tuned
krauss
parents: 33278
diff changeset
   451
            alt_name = Binding.empty,
bb65583ab70d absorbed inductive_wrap function into Function_Core; more conventional argument order; tuned
krauss
parents: 33278
diff changeset
   452
            coind = false,
bb65583ab70d absorbed inductive_wrap function into Function_Core; more conventional argument order; tuned
krauss
parents: 33278
diff changeset
   453
            no_elim = false,
bb65583ab70d absorbed inductive_wrap function into Function_Core; more conventional argument order; tuned
krauss
parents: 33278
diff changeset
   454
            no_ind = false,
49170
03bee3a6a1b7 discontinued obsolete fork_mono to loosen some brakes -- NB: TTY interaction has Goal.future_proofs disabled due to missing Future.worker_task;
wenzelm
parents: 46217
diff changeset
   455
            skip_mono = true}
33348
bb65583ab70d absorbed inductive_wrap function into Function_Core; more conventional argument order; tuned
krauss
parents: 33278
diff changeset
   456
          [binding] (* relation *)
bb65583ab70d absorbed inductive_wrap function into Function_Core; more conventional argument order; tuned
krauss
parents: 33278
diff changeset
   457
          [] (* no parameters *)
bb65583ab70d absorbed inductive_wrap function into Function_Core; more conventional argument order; tuned
krauss
parents: 33278
diff changeset
   458
          (map (fn t => (Attrib.empty_binding, t)) intrs) (* intro rules *)
bb65583ab70d absorbed inductive_wrap function into Function_Core; more conventional argument order; tuned
krauss
parents: 33278
diff changeset
   459
          [] (* no special monos *)
59880
30687c3f2b10 clarified role of naming for background theory: transform_binding (e.g. for "concealed" flag) uses naming of hypothetical context;
wenzelm
parents: 59859
diff changeset
   460
      ||> Proof_Context.restore_naming lthy
33348
bb65583ab70d absorbed inductive_wrap function into Function_Core; more conventional argument order; tuned
krauss
parents: 33278
diff changeset
   461
bb65583ab70d absorbed inductive_wrap function into Function_Core; more conventional argument order; tuned
krauss
parents: 33278
diff changeset
   462
    fun requantify orig_intro thm =
bb65583ab70d absorbed inductive_wrap function into Function_Core; more conventional argument order; tuned
krauss
parents: 33278
diff changeset
   463
      let
bb65583ab70d absorbed inductive_wrap function into Function_Core; more conventional argument order; tuned
krauss
parents: 33278
diff changeset
   464
        val (qs, t) = dest_all_all orig_intro
42483
39eefaef816a eliminated obsolete Function_Lib.frees_in_term;
wenzelm
parents: 42362
diff changeset
   465
        val frees = Variable.add_frees lthy t [] |> remove (op =) (Binding.name_of R, T)
59582
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59580
diff changeset
   466
        val vars = Term.add_vars (Thm.prop_of thm) []
33348
bb65583ab70d absorbed inductive_wrap function into Function_Core; more conventional argument order; tuned
krauss
parents: 33278
diff changeset
   467
        val varmap = AList.lookup (op =) (frees ~~ map fst vars)
42483
39eefaef816a eliminated obsolete Function_Lib.frees_in_term;
wenzelm
parents: 42362
diff changeset
   468
          #> the_default ("", 0)
33348
bb65583ab70d absorbed inductive_wrap function into Function_Core; more conventional argument order; tuned
krauss
parents: 33278
diff changeset
   469
      in
bb65583ab70d absorbed inductive_wrap function into Function_Core; more conventional argument order; tuned
krauss
parents: 33278
diff changeset
   470
        fold_rev (fn Free (n, T) =>
59621
291934bac95e Thm.cterm_of and Thm.ctyp_of operate on local context;
wenzelm
parents: 59618
diff changeset
   471
          forall_intr_rename (n, Thm.cterm_of lthy (Var (varmap (n, T), T)))) qs thm
33348
bb65583ab70d absorbed inductive_wrap function into Function_Core; more conventional argument order; tuned
krauss
parents: 33278
diff changeset
   472
      end
bb65583ab70d absorbed inductive_wrap function into Function_Core; more conventional argument order; tuned
krauss
parents: 33278
diff changeset
   473
  in
34232
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
   474
    ((Rdef, map2 requantify intrs intrs_gen, forall_intr_vars elim_gen, induct), lthy)
33348
bb65583ab70d absorbed inductive_wrap function into Function_Core; more conventional argument order; tuned
krauss
parents: 33278
diff changeset
   475
  end
bb65583ab70d absorbed inductive_wrap function into Function_Core; more conventional argument order; tuned
krauss
parents: 33278
diff changeset
   476
33099
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   477
fun define_graph Gname fvar domT ranT clauses RCss lthy =
33349
krauss
parents: 33348
diff changeset
   478
  let
krauss
parents: 33348
diff changeset
   479
    val GT = domT --> ranT --> boolT
krauss
parents: 33348
diff changeset
   480
    val (Gvar as (n, T)) = singleton (Variable.variant_frees lthy []) (Gname, GT)
33099
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   481
33349
krauss
parents: 33348
diff changeset
   482
    fun mk_GIntro (ClauseContext {qs, gs, lhs, rhs, ...}) RCs =
krauss
parents: 33348
diff changeset
   483
      let
krauss
parents: 33348
diff changeset
   484
        fun mk_h_assm (rcfix, rcassm, rcarg) =
krauss
parents: 33348
diff changeset
   485
          HOLogic.mk_Trueprop (Free Gvar $ rcarg $ (fvar $ rcarg))
59582
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59580
diff changeset
   486
          |> fold_rev (curry Logic.mk_implies o Thm.prop_of) rcassm
33349
krauss
parents: 33348
diff changeset
   487
          |> fold_rev (Logic.all o Free) rcfix
krauss
parents: 33348
diff changeset
   488
      in
krauss
parents: 33348
diff changeset
   489
        HOLogic.mk_Trueprop (Free Gvar $ lhs $ rhs)
krauss
parents: 33348
diff changeset
   490
        |> fold_rev (curry Logic.mk_implies o mk_h_assm) RCs
krauss
parents: 33348
diff changeset
   491
        |> fold_rev (curry Logic.mk_implies) gs
krauss
parents: 33348
diff changeset
   492
        |> fold_rev Logic.all (fvar :: qs)
krauss
parents: 33348
diff changeset
   493
      end
33099
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   494
33349
krauss
parents: 33348
diff changeset
   495
    val G_intros = map2 mk_GIntro clauses RCss
krauss
parents: 33348
diff changeset
   496
  in
krauss
parents: 33348
diff changeset
   497
    inductive_def ((Binding.name n, T), NoSyn) G_intros lthy
krauss
parents: 33348
diff changeset
   498
  end
33099
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   499
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   500
fun define_function fdefname (fname, mixfix) domT ranT G default lthy =
33349
krauss
parents: 33348
diff changeset
   501
  let
krauss
parents: 33348
diff changeset
   502
    val f_def =
55085
0e8e4dc55866 moved 'fundef_cong' attribute (and other basic 'fun' stuff) up the dependency chain
blanchet
parents: 53603
diff changeset
   503
      Abs ("x", domT, Const (@{const_name Fun_Def.THE_default}, ranT --> (ranT --> boolT) --> ranT)
33349
krauss
parents: 33348
diff changeset
   504
        $ (default $ Bound 0) $ Abs ("y", ranT, G $ Bound 1 $ Bound 0))
krauss
parents: 33348
diff changeset
   505
      |> Syntax.check_term lthy
61121
efe8b18306b7 do not expose low-level "_def" facts of 'function' definitions, to avoid potential confusion with the situation of plain 'definition';
wenzelm
parents: 60801
diff changeset
   506
    val def_binding =
efe8b18306b7 do not expose low-level "_def" facts of 'function' definitions, to avoid potential confusion with the situation of plain 'definition';
wenzelm
parents: 60801
diff changeset
   507
      if Config.get lthy function_defs then (Binding.name fdefname, [])
efe8b18306b7 do not expose low-level "_def" facts of 'function' definitions, to avoid potential confusion with the situation of plain 'definition';
wenzelm
parents: 60801
diff changeset
   508
      else Attrib.empty_binding
33349
krauss
parents: 33348
diff changeset
   509
  in
33766
c679f05600cd adapted Local_Theory.define -- eliminated odd thm kind;
wenzelm
parents: 33671
diff changeset
   510
    Local_Theory.define
61121
efe8b18306b7 do not expose low-level "_def" facts of 'function' definitions, to avoid potential confusion with the situation of plain 'definition';
wenzelm
parents: 60801
diff changeset
   511
      ((Binding.name (function_name fname), mixfix), (def_binding, f_def)) lthy
33349
krauss
parents: 33348
diff changeset
   512
  end
33099
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   513
33855
cd8acf137c9c eliminated dead code and some unused bindings, reported by polyml
krauss
parents: 33766
diff changeset
   514
fun define_recursion_relation Rname domT qglrs clauses RCss lthy =
33349
krauss
parents: 33348
diff changeset
   515
  let
krauss
parents: 33348
diff changeset
   516
    val RT = domT --> domT --> boolT
krauss
parents: 33348
diff changeset
   517
    val (Rvar as (n, T)) = singleton (Variable.variant_frees lthy []) (Rname, RT)
33099
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   518
33349
krauss
parents: 33348
diff changeset
   519
    fun mk_RIntro (ClauseContext {qs, gs, lhs, ...}, (oqs, _, _, _)) (rcfix, rcassm, rcarg) =
krauss
parents: 33348
diff changeset
   520
      HOLogic.mk_Trueprop (Free Rvar $ rcarg $ lhs)
59582
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59580
diff changeset
   521
      |> fold_rev (curry Logic.mk_implies o Thm.prop_of) rcassm
33349
krauss
parents: 33348
diff changeset
   522
      |> fold_rev (curry Logic.mk_implies) gs
krauss
parents: 33348
diff changeset
   523
      |> fold_rev (Logic.all o Free) rcfix
krauss
parents: 33348
diff changeset
   524
      |> fold_rev mk_forall_rename (map fst oqs ~~ qs)
krauss
parents: 33348
diff changeset
   525
      (* "!!qs xs. CS ==> G => (r, lhs) : R" *)
33099
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   526
33349
krauss
parents: 33348
diff changeset
   527
    val R_intross = map2 (map o mk_RIntro) (clauses ~~ qglrs) RCss
33099
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   528
33349
krauss
parents: 33348
diff changeset
   529
    val ((R, RIntro_thms, R_elim, _), lthy) =
krauss
parents: 33348
diff changeset
   530
      inductive_def ((Binding.name n, T), NoSyn) (flat R_intross) lthy
krauss
parents: 33348
diff changeset
   531
  in
krauss
parents: 33348
diff changeset
   532
    ((R, Library.unflat R_intross RIntro_thms, R_elim), lthy)
krauss
parents: 33348
diff changeset
   533
  end
33099
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   534
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   535
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   536
fun fix_globals domT ranT fvar ctxt =
34232
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
   537
  let
59618
e6939796717e clarified context;
wenzelm
parents: 59582
diff changeset
   538
    val ([h, y, x, z, a, D, P, Pbool], ctxt') = Variable.variant_fixes
34232
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
   539
      ["h_fd", "y_fd", "x_fd", "z_fd", "a_fd", "D_fd", "P_fd", "Pb_fd"] ctxt
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
   540
  in
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
   541
    (Globals {h = Free (h, domT --> ranT),
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
   542
      y = Free (y, ranT),
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
   543
      x = Free (x, domT),
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
   544
      z = Free (z, domT),
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
   545
      a = Free (a, domT),
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
   546
      D = Free (D, domT --> boolT),
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
   547
      P = Free (P, domT --> boolT),
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
   548
      Pbool = Free (Pbool, boolT),
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
   549
      fvar = fvar,
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
   550
      domT = domT,
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
   551
      ranT = ranT},
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
   552
    ctxt')
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
   553
  end
33099
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   554
59627
bb1e4a35d506 clarified context;
wenzelm
parents: 59621
diff changeset
   555
fun inst_RC ctxt fvar f (rcfix, rcassm, rcarg) =
34232
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
   556
  let
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
   557
    fun inst_term t = subst_bound(f, abstract_over (fvar, t))
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
   558
  in
59627
bb1e4a35d506 clarified context;
wenzelm
parents: 59621
diff changeset
   559
    (rcfix, map (Thm.assume o Thm.cterm_of ctxt o inst_term o Thm.prop_of) rcassm, inst_term rcarg)
34232
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
   560
  end
33099
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   561
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   562
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   563
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   564
(**********************************************************
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   565
 *                   PROVING THE RULES
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   566
 **********************************************************)
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   567
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 49170
diff changeset
   568
fun mk_psimps ctxt globals R clauses valthms f_iff graph_is_function =
34232
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
   569
  let
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
   570
    val Globals {domT, z, ...} = globals
33099
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   571
59582
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59580
diff changeset
   572
    fun mk_psimp
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59580
diff changeset
   573
      (ClauseInfo {qglr = (oqs, _, _, _), cdata = ClauseContext {cqs, lhs, ags, ...}, ...}) valthm =
34232
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
   574
      let
59618
e6939796717e clarified context;
wenzelm
parents: 59582
diff changeset
   575
        val lhs_acc =
59621
291934bac95e Thm.cterm_of and Thm.ctyp_of operate on local context;
wenzelm
parents: 59618
diff changeset
   576
          Thm.cterm_of ctxt (HOLogic.mk_Trueprop (mk_acc domT R $ lhs)) (* "acc R lhs" *)
59618
e6939796717e clarified context;
wenzelm
parents: 59582
diff changeset
   577
        val z_smaller =
59621
291934bac95e Thm.cterm_of and Thm.ctyp_of operate on local context;
wenzelm
parents: 59618
diff changeset
   578
          Thm.cterm_of ctxt (HOLogic.mk_Trueprop (R $ z $ lhs)) (* "R z lhs" *)
34232
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
   579
      in
36945
9bec62c10714 less pervasive names from structure Thm;
wenzelm
parents: 36936
diff changeset
   580
        ((Thm.assume z_smaller) RS ((Thm.assume lhs_acc) RS acc_downward))
34232
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
   581
        |> (fn it => it COMP graph_is_function)
36945
9bec62c10714 less pervasive names from structure Thm;
wenzelm
parents: 36936
diff changeset
   582
        |> Thm.implies_intr z_smaller
59621
291934bac95e Thm.cterm_of and Thm.ctyp_of operate on local context;
wenzelm
parents: 59618
diff changeset
   583
        |> Thm.forall_intr (Thm.cterm_of ctxt  z)
34232
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
   584
        |> (fn it => it COMP valthm)
36945
9bec62c10714 less pervasive names from structure Thm;
wenzelm
parents: 36936
diff changeset
   585
        |> Thm.implies_intr lhs_acc
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 49170
diff changeset
   586
        |> asm_simplify (put_simpset HOL_basic_ss ctxt addsimps [f_iff])
59582
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59580
diff changeset
   587
        |> fold_rev (Thm.implies_intr o Thm.cprop_of) ags
34232
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
   588
        |> fold_rev forall_intr_rename (map fst oqs ~~ cqs)
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
   589
      end
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
   590
  in
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
   591
    map2 mk_psimp clauses valthms
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
   592
  end
33099
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   593
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   594
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   595
(** Induction rule **)
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   596
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   597
34065
6f8f9835e219 moved predicate rules to Predicate.thy; weakened default dest rule predicate1D (is not that reliable wrt. sets)
haftmann
parents: 33855
diff changeset
   598
val acc_subset_induct = @{thm predicate1I} RS @{thm accp_subset_induct}
33099
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   599
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   600
59627
bb1e4a35d506 clarified context;
wenzelm
parents: 59621
diff changeset
   601
fun mk_partial_induct_rule ctxt globals R complete_thm clauses =
34232
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
   602
  let
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
   603
    val Globals {domT, x, z, a, P, D, ...} = globals
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
   604
    val acc_R = mk_acc domT R
33099
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   605
59627
bb1e4a35d506 clarified context;
wenzelm
parents: 59621
diff changeset
   606
    val x_D = Thm.assume (Thm.cterm_of ctxt (HOLogic.mk_Trueprop (D $ x)))
bb1e4a35d506 clarified context;
wenzelm
parents: 59621
diff changeset
   607
    val a_D = Thm.cterm_of ctxt (HOLogic.mk_Trueprop (D $ a))
33099
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   608
59627
bb1e4a35d506 clarified context;
wenzelm
parents: 59621
diff changeset
   609
    val D_subset =
bb1e4a35d506 clarified context;
wenzelm
parents: 59621
diff changeset
   610
      Thm.cterm_of ctxt (Logic.all x
bb1e4a35d506 clarified context;
wenzelm
parents: 59621
diff changeset
   611
        (Logic.mk_implies (HOLogic.mk_Trueprop (D $ x), HOLogic.mk_Trueprop (acc_R $ x))))
33099
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   612
34232
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
   613
    val D_dcl = (* "!!x z. [| x: D; (z,x):R |] ==> z:D" *)
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
   614
      Logic.all x (Logic.all z (Logic.mk_implies (HOLogic.mk_Trueprop (D $ x),
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
   615
        Logic.mk_implies (HOLogic.mk_Trueprop (R $ z $ x),
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
   616
          HOLogic.mk_Trueprop (D $ z)))))
59627
bb1e4a35d506 clarified context;
wenzelm
parents: 59621
diff changeset
   617
      |> Thm.cterm_of ctxt
33099
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   618
34232
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
   619
    (* Inductive Hypothesis: !!z. (z,x):R ==> P z *)
46217
7b19666f0e3d renamed Term.all to Logic.all_const, in accordance to HOLogic.all_const;
wenzelm
parents: 42793
diff changeset
   620
    val ihyp = Logic.all_const domT $ Abs ("z", domT,
34232
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
   621
      Logic.mk_implies (HOLogic.mk_Trueprop (R $ Bound 0 $ x),
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
   622
        HOLogic.mk_Trueprop (P $ Bound 0)))
59627
bb1e4a35d506 clarified context;
wenzelm
parents: 59621
diff changeset
   623
      |> Thm.cterm_of ctxt
33099
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   624
36945
9bec62c10714 less pervasive names from structure Thm;
wenzelm
parents: 36936
diff changeset
   625
    val aihyp = Thm.assume ihyp
33099
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   626
34232
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
   627
    fun prove_case clause =
33099
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   628
      let
59627
bb1e4a35d506 clarified context;
wenzelm
parents: 59621
diff changeset
   629
        val ClauseInfo {cdata = ClauseContext {ctxt = ctxt1, qs, cqs, ags, gs, lhs, case_hyp, ...},
34232
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
   630
          RCs, qglr = (oqs, _, _, _), ...} = clause
33099
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   631
34232
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
   632
        val case_hyp_conv = K (case_hyp RS eq_reflection)
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
   633
        local open Conv in
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
   634
          val lhs_D = fconv_rule (arg_conv (arg_conv (case_hyp_conv))) x_D
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
   635
          val sih =
36936
c52d1c130898 incorporated further conversions and conversionals, after some minor tuning;
wenzelm
parents: 36270
diff changeset
   636
            fconv_rule (Conv.binder_conv
59627
bb1e4a35d506 clarified context;
wenzelm
parents: 59621
diff changeset
   637
              (K (arg1_conv (arg_conv (arg_conv case_hyp_conv)))) ctxt1) aihyp
34232
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
   638
        end
33099
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   639
34232
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
   640
        fun mk_Prec (RCInfo {llRI, RIvs, CCas, rcarg, ...}) = sih
59627
bb1e4a35d506 clarified context;
wenzelm
parents: 59621
diff changeset
   641
          |> Thm.forall_elim (Thm.cterm_of ctxt rcarg)
34232
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
   642
          |> Thm.elim_implies llRI
59582
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59580
diff changeset
   643
          |> fold_rev (Thm.implies_intr o Thm.cprop_of) CCas
59627
bb1e4a35d506 clarified context;
wenzelm
parents: 59621
diff changeset
   644
          |> fold_rev (Thm.forall_intr o Thm.cterm_of ctxt o Free) RIvs
33099
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   645
34232
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
   646
        val P_recs = map mk_Prec RCs   (*  [P rec1, P rec2, ... ]  *)
33099
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   647
34232
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
   648
        val step = HOLogic.mk_Trueprop (P $ lhs)
59582
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59580
diff changeset
   649
          |> fold_rev (curry Logic.mk_implies o Thm.prop_of) P_recs
34232
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
   650
          |> fold_rev (curry Logic.mk_implies) gs
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
   651
          |> curry Logic.mk_implies (HOLogic.mk_Trueprop (D $ lhs))
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
   652
          |> fold_rev mk_forall_rename (map fst oqs ~~ qs)
59627
bb1e4a35d506 clarified context;
wenzelm
parents: 59621
diff changeset
   653
          |> Thm.cterm_of ctxt
33099
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   654
36945
9bec62c10714 less pervasive names from structure Thm;
wenzelm
parents: 36936
diff changeset
   655
        val P_lhs = Thm.assume step
9bec62c10714 less pervasive names from structure Thm;
wenzelm
parents: 36936
diff changeset
   656
          |> fold Thm.forall_elim cqs
34232
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
   657
          |> Thm.elim_implies lhs_D
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
   658
          |> fold Thm.elim_implies ags
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
   659
          |> fold Thm.elim_implies P_recs
33099
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   660
59627
bb1e4a35d506 clarified context;
wenzelm
parents: 59621
diff changeset
   661
        val res = Thm.cterm_of ctxt (HOLogic.mk_Trueprop (P $ x))
34232
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
   662
          |> Conv.arg_conv (Conv.arg_conv case_hyp_conv)
36945
9bec62c10714 less pervasive names from structure Thm;
wenzelm
parents: 36936
diff changeset
   663
          |> Thm.symmetric (* P lhs == P x *)
9bec62c10714 less pervasive names from structure Thm;
wenzelm
parents: 36936
diff changeset
   664
          |> (fn eql => Thm.equal_elim eql P_lhs) (* "P x" *)
59582
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59580
diff changeset
   665
          |> Thm.implies_intr (Thm.cprop_of case_hyp)
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59580
diff changeset
   666
          |> fold_rev (Thm.implies_intr o Thm.cprop_of) ags
36945
9bec62c10714 less pervasive names from structure Thm;
wenzelm
parents: 36936
diff changeset
   667
          |> fold_rev Thm.forall_intr cqs
33099
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   668
      in
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   669
        (res, step)
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   670
      end
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   671
34232
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
   672
    val (cases, steps) = split_list (map prove_case clauses)
33099
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   673
34232
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
   674
    val istep = complete_thm
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
   675
      |> Thm.forall_elim_vars 0
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
   676
      |> fold (curry op COMP) cases (*  P x  *)
36945
9bec62c10714 less pervasive names from structure Thm;
wenzelm
parents: 36936
diff changeset
   677
      |> Thm.implies_intr ihyp
59582
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59580
diff changeset
   678
      |> Thm.implies_intr (Thm.cprop_of x_D)
59627
bb1e4a35d506 clarified context;
wenzelm
parents: 59621
diff changeset
   679
      |> Thm.forall_intr (Thm.cterm_of ctxt x)
33099
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   680
34232
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
   681
    val subset_induct_rule =
33099
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   682
      acc_subset_induct
36945
9bec62c10714 less pervasive names from structure Thm;
wenzelm
parents: 36936
diff changeset
   683
      |> (curry op COMP) (Thm.assume D_subset)
9bec62c10714 less pervasive names from structure Thm;
wenzelm
parents: 36936
diff changeset
   684
      |> (curry op COMP) (Thm.assume D_dcl)
9bec62c10714 less pervasive names from structure Thm;
wenzelm
parents: 36936
diff changeset
   685
      |> (curry op COMP) (Thm.assume a_D)
34232
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
   686
      |> (curry op COMP) istep
36945
9bec62c10714 less pervasive names from structure Thm;
wenzelm
parents: 36936
diff changeset
   687
      |> fold_rev Thm.implies_intr steps
9bec62c10714 less pervasive names from structure Thm;
wenzelm
parents: 36936
diff changeset
   688
      |> Thm.implies_intr a_D
9bec62c10714 less pervasive names from structure Thm;
wenzelm
parents: 36936
diff changeset
   689
      |> Thm.implies_intr D_dcl
9bec62c10714 less pervasive names from structure Thm;
wenzelm
parents: 36936
diff changeset
   690
      |> Thm.implies_intr D_subset
33099
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   691
34232
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
   692
    val simple_induct_rule =
33099
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   693
      subset_induct_rule
59627
bb1e4a35d506 clarified context;
wenzelm
parents: 59621
diff changeset
   694
      |> Thm.forall_intr (Thm.cterm_of ctxt D)
bb1e4a35d506 clarified context;
wenzelm
parents: 59621
diff changeset
   695
      |> Thm.forall_elim (Thm.cterm_of ctxt acc_R)
60752
b48830b670a1 prefer tactics with explicit context;
wenzelm
parents: 60643
diff changeset
   696
      |> assume_tac ctxt 1 |> Seq.hd
34232
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
   697
      |> (curry op COMP) (acc_downward
60801
7664e0916eec tuned signature;
wenzelm
parents: 60752
diff changeset
   698
        |> (Thm.instantiate' [SOME (Thm.ctyp_of ctxt domT)] (map (SOME o Thm.cterm_of ctxt) [R, x, z]))
59627
bb1e4a35d506 clarified context;
wenzelm
parents: 59621
diff changeset
   699
        |> Thm.forall_intr (Thm.cterm_of ctxt z)
bb1e4a35d506 clarified context;
wenzelm
parents: 59621
diff changeset
   700
        |> Thm.forall_intr (Thm.cterm_of ctxt x))
bb1e4a35d506 clarified context;
wenzelm
parents: 59621
diff changeset
   701
      |> Thm.forall_intr (Thm.cterm_of ctxt a)
bb1e4a35d506 clarified context;
wenzelm
parents: 59621
diff changeset
   702
      |> Thm.forall_intr (Thm.cterm_of ctxt P)
34232
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
   703
  in
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
   704
    simple_induct_rule
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
   705
  end
33099
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   706
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   707
34232
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
   708
(* FIXME: broken by design *)
33099
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   709
fun mk_domain_intro ctxt (Globals {domT, ...}) R R_cases clause =
34232
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
   710
  let
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
   711
    val ClauseInfo {cdata = ClauseContext {gs, lhs, cqs, ...},
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
   712
      qglr = (oqs, _, _, _), ...} = clause
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
   713
    val goal = HOLogic.mk_Trueprop (mk_acc domT R $ lhs)
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
   714
      |> fold_rev (curry Logic.mk_implies) gs
59621
291934bac95e Thm.cterm_of and Thm.ctyp_of operate on local context;
wenzelm
parents: 59618
diff changeset
   715
      |> Thm.cterm_of ctxt
34232
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
   716
  in
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
   717
    Goal.init goal
59498
50b60f501b05 proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents: 58963
diff changeset
   718
    |> (SINGLE (resolve_tac ctxt [accI] 1)) |> the
50b60f501b05 proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents: 58963
diff changeset
   719
    |> (SINGLE (eresolve_tac ctxt [Thm.forall_elim_vars 0 R_cases] 1))  |> the
42793
88bee9f6eec7 proper Proof.context for classical tactics;
wenzelm
parents: 42483
diff changeset
   720
    |> (SINGLE (auto_tac ctxt)) |> the
34232
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
   721
    |> Goal.conclude
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
   722
    |> fold_rev forall_intr_rename (map fst oqs ~~ cqs)
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
   723
  end
33099
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   724
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   725
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   726
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   727
(** Termination rule **)
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   728
34232
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
   729
val wf_induct_rule = @{thm Wellfounded.wfP_induct_rule}
55085
0e8e4dc55866 moved 'fundef_cong' attribute (and other basic 'fun' stuff) up the dependency chain
blanchet
parents: 53603
diff changeset
   730
val wf_in_rel = @{thm Fun_Def.wf_in_rel}
0e8e4dc55866 moved 'fundef_cong' attribute (and other basic 'fun' stuff) up the dependency chain
blanchet
parents: 53603
diff changeset
   731
val in_rel_def = @{thm Fun_Def.in_rel_def}
33099
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   732
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 49170
diff changeset
   733
fun mk_nest_term_case ctxt globals R' ihyp clause =
34232
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
   734
  let
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
   735
    val Globals {z, ...} = globals
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
   736
    val ClauseInfo {cdata = ClauseContext {qs, cqs, ags, lhs, case_hyp, ...}, tree,
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
   737
      qglr=(oqs, _, _, _), ...} = clause
33099
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   738
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 49170
diff changeset
   739
    val ih_case = full_simplify (put_simpset HOL_basic_ss ctxt addsimps [case_hyp]) ihyp
33099
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   740
34232
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
   741
    fun step (fixes, assumes) (_ $ arg) u (sub,(hyps,thms)) =
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
   742
      let
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
   743
        val used = (u @ sub)
59618
e6939796717e clarified context;
wenzelm
parents: 59582
diff changeset
   744
          |> map (fn (ctx, thm) => Function_Context_Tree.export_thm ctxt ctx thm)
33099
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   745
34232
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
   746
        val hyp = HOLogic.mk_Trueprop (R' $ arg $ lhs)
59582
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59580
diff changeset
   747
          |> fold_rev (curry Logic.mk_implies o Thm.prop_of) used (* additional hyps *)
58816
aab139c0003f modernized setup;
wenzelm
parents: 58634
diff changeset
   748
          |> Function_Context_Tree.export_term (fixes, assumes)
59582
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59580
diff changeset
   749
          |> fold_rev (curry Logic.mk_implies o Thm.prop_of) ags
34232
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
   750
          |> fold_rev mk_forall_rename (map fst oqs ~~ qs)
59621
291934bac95e Thm.cterm_of and Thm.ctyp_of operate on local context;
wenzelm
parents: 59618
diff changeset
   751
          |> Thm.cterm_of ctxt
33099
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   752
36945
9bec62c10714 less pervasive names from structure Thm;
wenzelm
parents: 36936
diff changeset
   753
        val thm = Thm.assume hyp
9bec62c10714 less pervasive names from structure Thm;
wenzelm
parents: 36936
diff changeset
   754
          |> fold Thm.forall_elim cqs
34232
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
   755
          |> fold Thm.elim_implies ags
59618
e6939796717e clarified context;
wenzelm
parents: 59582
diff changeset
   756
          |> Function_Context_Tree.import_thm ctxt (fixes, assumes)
34232
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
   757
          |> fold Thm.elim_implies used (*  "(arg, lhs) : R'"  *)
33099
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   758
34232
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
   759
        val z_eq_arg = HOLogic.mk_Trueprop (mk_eq (z, arg))
59621
291934bac95e Thm.cterm_of and Thm.ctyp_of operate on local context;
wenzelm
parents: 59618
diff changeset
   760
          |> Thm.cterm_of ctxt |> Thm.assume
33099
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   761
34232
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
   762
        val acc = thm COMP ih_case
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
   763
        val z_acc_local = acc
36945
9bec62c10714 less pervasive names from structure Thm;
wenzelm
parents: 36936
diff changeset
   764
          |> Conv.fconv_rule
9bec62c10714 less pervasive names from structure Thm;
wenzelm
parents: 36936
diff changeset
   765
              (Conv.arg_conv (Conv.arg_conv (K (Thm.symmetric (z_eq_arg RS eq_reflection)))))
33099
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   766
34232
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
   767
        val ethm = z_acc_local
59618
e6939796717e clarified context;
wenzelm
parents: 59582
diff changeset
   768
          |> Function_Context_Tree.export_thm ctxt (fixes,
34232
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
   769
               z_eq_arg :: case_hyp :: ags @ assumes)
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
   770
          |> fold_rev forall_intr_rename (map fst oqs ~~ cqs)
33099
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   771
34232
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
   772
        val sub' = sub @ [(([],[]), acc)]
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
   773
      in
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
   774
        (sub', (hyp :: hyps, ethm :: thms))
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
   775
      end
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
   776
      | step _ _ _ _ = raise Match
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
   777
  in
58816
aab139c0003f modernized setup;
wenzelm
parents: 58634
diff changeset
   778
    Function_Context_Tree.traverse_tree step tree
34232
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
   779
  end
33099
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   780
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   781
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 49170
diff changeset
   782
fun mk_nest_term_rule ctxt globals R R_cases clauses =
34232
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
   783
  let
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
   784
    val Globals { domT, x, z, ... } = globals
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
   785
    val acc_R = mk_acc domT R
33099
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   786
34232
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
   787
    val R' = Free ("R", fastype_of R)
33099
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   788
34232
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
   789
    val Rrel = Free ("R", HOLogic.mk_setT (HOLogic.mk_prodT (domT, domT)))
55085
0e8e4dc55866 moved 'fundef_cong' attribute (and other basic 'fun' stuff) up the dependency chain
blanchet
parents: 53603
diff changeset
   790
    val inrel_R = Const (@{const_name Fun_Def.in_rel},
34232
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
   791
      HOLogic.mk_setT (HOLogic.mk_prodT (domT, domT)) --> fastype_of R) $ Rrel
33099
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   792
34232
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
   793
    val wfR' = HOLogic.mk_Trueprop (Const (@{const_name Wellfounded.wfP},
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
   794
      (domT --> domT --> boolT) --> boolT) $ R')
59621
291934bac95e Thm.cterm_of and Thm.ctyp_of operate on local context;
wenzelm
parents: 59618
diff changeset
   795
      |> Thm.cterm_of ctxt (* "wf R'" *)
33099
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   796
34232
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
   797
    (* Inductive Hypothesis: !!z. (z,x):R' ==> z : acc R *)
46217
7b19666f0e3d renamed Term.all to Logic.all_const, in accordance to HOLogic.all_const;
wenzelm
parents: 42793
diff changeset
   798
    val ihyp = Logic.all_const domT $ Abs ("z", domT,
34232
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
   799
      Logic.mk_implies (HOLogic.mk_Trueprop (R' $ Bound 0 $ x),
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
   800
        HOLogic.mk_Trueprop (acc_R $ Bound 0)))
59621
291934bac95e Thm.cterm_of and Thm.ctyp_of operate on local context;
wenzelm
parents: 59618
diff changeset
   801
      |> Thm.cterm_of ctxt
33099
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   802
36945
9bec62c10714 less pervasive names from structure Thm;
wenzelm
parents: 36936
diff changeset
   803
    val ihyp_a = Thm.assume ihyp |> Thm.forall_elim_vars 0
33099
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   804
59621
291934bac95e Thm.cterm_of and Thm.ctyp_of operate on local context;
wenzelm
parents: 59618
diff changeset
   805
    val R_z_x = Thm.cterm_of ctxt (HOLogic.mk_Trueprop (R $ z $ x))
33099
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   806
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 49170
diff changeset
   807
    val (hyps, cases) = fold (mk_nest_term_case ctxt globals R' ihyp_a) clauses ([], [])
34232
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
   808
  in
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
   809
    R_cases
59621
291934bac95e Thm.cterm_of and Thm.ctyp_of operate on local context;
wenzelm
parents: 59618
diff changeset
   810
    |> Thm.forall_elim (Thm.cterm_of ctxt z)
291934bac95e Thm.cterm_of and Thm.ctyp_of operate on local context;
wenzelm
parents: 59618
diff changeset
   811
    |> Thm.forall_elim (Thm.cterm_of ctxt x)
291934bac95e Thm.cterm_of and Thm.ctyp_of operate on local context;
wenzelm
parents: 59618
diff changeset
   812
    |> Thm.forall_elim (Thm.cterm_of ctxt (acc_R $ z))
36945
9bec62c10714 less pervasive names from structure Thm;
wenzelm
parents: 36936
diff changeset
   813
    |> curry op COMP (Thm.assume R_z_x)
34232
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
   814
    |> fold_rev (curry op COMP) cases
36945
9bec62c10714 less pervasive names from structure Thm;
wenzelm
parents: 36936
diff changeset
   815
    |> Thm.implies_intr R_z_x
59621
291934bac95e Thm.cterm_of and Thm.ctyp_of operate on local context;
wenzelm
parents: 59618
diff changeset
   816
    |> Thm.forall_intr (Thm.cterm_of ctxt z)
34232
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
   817
    |> (fn it => it COMP accI)
36945
9bec62c10714 less pervasive names from structure Thm;
wenzelm
parents: 36936
diff changeset
   818
    |> Thm.implies_intr ihyp
59621
291934bac95e Thm.cterm_of and Thm.ctyp_of operate on local context;
wenzelm
parents: 59618
diff changeset
   819
    |> Thm.forall_intr (Thm.cterm_of ctxt x)
52467
24c6ddb48cb8 tuned signature;
wenzelm
parents: 52456
diff changeset
   820
    |> (fn it => Drule.compose (it, 2, wf_induct_rule))
36945
9bec62c10714 less pervasive names from structure Thm;
wenzelm
parents: 36936
diff changeset
   821
    |> curry op RS (Thm.assume wfR')
34232
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
   822
    |> forall_intr_vars
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
   823
    |> (fn it => it COMP allI)
36945
9bec62c10714 less pervasive names from structure Thm;
wenzelm
parents: 36936
diff changeset
   824
    |> fold Thm.implies_intr hyps
9bec62c10714 less pervasive names from structure Thm;
wenzelm
parents: 36936
diff changeset
   825
    |> Thm.implies_intr wfR'
59621
291934bac95e Thm.cterm_of and Thm.ctyp_of operate on local context;
wenzelm
parents: 59618
diff changeset
   826
    |> Thm.forall_intr (Thm.cterm_of ctxt R')
291934bac95e Thm.cterm_of and Thm.ctyp_of operate on local context;
wenzelm
parents: 59618
diff changeset
   827
    |> Thm.forall_elim (Thm.cterm_of ctxt (inrel_R))
34232
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
   828
    |> curry op RS wf_in_rel
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 49170
diff changeset
   829
    |> full_simplify (put_simpset HOL_basic_ss ctxt addsimps [in_rel_def])
59621
291934bac95e Thm.cterm_of and Thm.ctyp_of operate on local context;
wenzelm
parents: 59618
diff changeset
   830
    |> Thm.forall_intr (Thm.cterm_of ctxt Rrel)
34232
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
   831
  end
33099
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   832
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   833
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   834
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   835
fun prepare_function config defname [((fname, fT), mixfix)] abstract_qglrs lthy =
34232
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
   836
  let
41846
b368a7aee46a removed support for tail-recursion from function package (now implemented by partial_function)
krauss
parents: 41417
diff changeset
   837
    val FunctionConfig {domintros, default=default_opt, ...} = config
33099
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   838
55396
91bc9f69a958 got rid of dynamic scoping the easy way
blanchet
parents: 55085
diff changeset
   839
    val default_str = the_default "%x. HOL.undefined" default_opt
34232
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
   840
    val fvar = Free (fname, fT)
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
   841
    val domT = domain_type fT
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
   842
    val ranT = range_type fT
33099
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   843
34232
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
   844
    val default = Syntax.parse_term lthy default_str
39288
f1ae2493d93f eliminated aliases of Type.constraint;
wenzelm
parents: 38920
diff changeset
   845
      |> Type.constraint fT |> Syntax.check_term lthy
34232
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
   846
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
   847
    val (globals, ctxt') = fix_globals domT ranT fvar lthy
33099
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   848
34232
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
   849
    val Globals { x, h, ... } = globals
33099
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   850
34232
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
   851
    val clauses = map (mk_clause_context x ctxt') abstract_qglrs
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
   852
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
   853
    val n = length abstract_qglrs
33099
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   854
34232
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
   855
    fun build_tree (ClauseContext { ctxt, rhs, ...}) =
58816
aab139c0003f modernized setup;
wenzelm
parents: 58634
diff changeset
   856
       Function_Context_Tree.mk_tree (fname, fT) h ctxt rhs
33099
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   857
34232
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
   858
    val trees = map build_tree clauses
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
   859
    val RCss = map find_calls trees
33099
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   860
34232
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
   861
    val ((G, GIntro_thms, G_elim, G_induct), lthy) =
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
   862
      PROFILE "def_graph" (define_graph (graph_name defname) fvar domT ranT clauses RCss) lthy
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
   863
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
   864
    val ((f, (_, f_defthm)), lthy) =
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
   865
      PROFILE "def_fun" (define_function (defname ^ "_sumC_def") (fname, mixfix) domT ranT G default) lthy
33099
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   866
59627
bb1e4a35d506 clarified context;
wenzelm
parents: 59621
diff changeset
   867
    val RCss = map (map (inst_RC lthy fvar f)) RCss
59618
e6939796717e clarified context;
wenzelm
parents: 59582
diff changeset
   868
    val trees = map (Function_Context_Tree.inst_tree lthy fvar f) trees
33099
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   869
34232
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
   870
    val ((R, RIntro_thmss, R_elim), lthy) =
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
   871
      PROFILE "def_rel" (define_recursion_relation (rel_name defname) domT abstract_qglrs clauses RCss) lthy
33099
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   872
52384
80c00a851de5 export dom predicate in the info record
krauss
parents: 52223
diff changeset
   873
    val dom = mk_acc domT R
34232
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
   874
    val (_, lthy) =
52384
80c00a851de5 export dom predicate in the info record
krauss
parents: 52223
diff changeset
   875
      Local_Theory.abbrev Syntax.mode_default ((Binding.name (dom_name defname), NoSyn), dom) lthy
33099
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   876
42361
23f352990944 modernized structure Proof_Context;
wenzelm
parents: 41846
diff changeset
   877
    val newthy = Proof_Context.theory_of lthy
34232
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
   878
    val clauses = map (transfer_clause_ctx newthy) clauses
33099
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   879
34232
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
   880
    val xclauses = PROFILE "xclauses"
58634
9f10d82e8188 added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
wenzelm
parents: 55990
diff changeset
   881
      (@{map 7} (mk_clause_info globals G f) (1 upto n) clauses abstract_qglrs trees
34232
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
   882
        RCss GIntro_thms) RIntro_thmss
33099
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   883
34232
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
   884
    val complete =
59621
291934bac95e Thm.cterm_of and Thm.ctyp_of operate on local context;
wenzelm
parents: 59618
diff changeset
   885
      mk_completeness globals clauses abstract_qglrs |> Thm.cterm_of lthy |> Thm.assume
33099
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   886
34232
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
   887
    val compat =
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
   888
      mk_compat_proof_obligations domT ranT fvar f abstract_qglrs
59621
291934bac95e Thm.cterm_of and Thm.ctyp_of operate on local context;
wenzelm
parents: 59618
diff changeset
   889
      |> map (Thm.cterm_of lthy #> Thm.assume)
33099
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   890
34232
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
   891
    val compat_store = store_compat_thms n compat
33099
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   892
34232
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
   893
    val (goalstate, values) = PROFILE "prove_stuff"
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
   894
      (prove_stuff lthy globals G f R xclauses complete compat
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
   895
         compat_store G_elim) f_defthm
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
   896
60643
9173467ec5b6 clarified context;
wenzelm
parents: 59880
diff changeset
   897
    fun mk_partial_rules newctxt provedgoal =
34232
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
   898
      let
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
   899
        val (graph_is_function, complete_thm) =
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
   900
          provedgoal
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
   901
          |> Conjunction.elim
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
   902
          |> apfst (Thm.forall_elim_vars 0)
33099
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   903
34232
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
   904
        val f_iff = graph_is_function RS (f_defthm RS ex1_implies_iff)
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
   905
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
   906
        val psimps = PROFILE "Proving simplification rules"
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 49170
diff changeset
   907
          (mk_psimps newctxt globals R xclauses values f_iff) graph_is_function
33099
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   908
34232
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
   909
        val simple_pinduct = PROFILE "Proving partial induction rule"
59627
bb1e4a35d506 clarified context;
wenzelm
parents: 59621
diff changeset
   910
          (mk_partial_induct_rule newctxt globals R complete_thm) xclauses
33099
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   911
34232
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
   912
        val total_intro = PROFILE "Proving nested termination rule"
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 49170
diff changeset
   913
          (mk_nest_term_rule newctxt globals R R_elim) xclauses
33099
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   914
34232
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
   915
        val dom_intros =
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
   916
          if domintros then SOME (PROFILE "Proving domain introduction rules"
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
   917
             (map (mk_domain_intro lthy globals R R_elim)) xclauses)
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
   918
           else NONE
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
   919
      in
53603
59ef06cda7b9 generate elim rules for elimination of function equalities;
Manuel Eberl
parents: 52467
diff changeset
   920
        FunctionResult {fs=[f], G=G, R=R, dom=dom,
59ef06cda7b9 generate elim rules for elimination of function equalities;
Manuel Eberl
parents: 52467
diff changeset
   921
          cases=[complete_thm], psimps=psimps, pelims=[],
59ef06cda7b9 generate elim rules for elimination of function equalities;
Manuel Eberl
parents: 52467
diff changeset
   922
          simple_pinducts=[simple_pinduct],
41846
b368a7aee46a removed support for tail-recursion from function package (now implemented by partial_function)
krauss
parents: 41417
diff changeset
   923
          termination=total_intro, domintros=dom_intros}
34232
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
   924
      end
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
   925
  in
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
   926
    ((f, goalstate, mk_partial_rules), lthy)
36a2a3029fd3 new year's resolution: reindented code in function package
krauss
parents: 34065
diff changeset
   927
  end
33099
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   928
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   929
b8cdd3d73022 function package: more standard names for structures and files
krauss
parents:
diff changeset
   930
end