src/HOL/Tools/Transfer/transfer.ML
author haftmann
Fri, 14 Jun 2019 08:34:27 +0000
changeset 70345 80a1aa30e24d
parent 70316 c61b7af108a6
child 70473 9179e7a98196
permissions -rw-r--r--
avoid pseudo-collection to be used in generated proofs
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
56524
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents: 56520
diff changeset
     1
(*  Title:      HOL/Tools/Transfer/transfer.ML
47325
ec6187036495 new transfer proof method
huffman
parents:
diff changeset
     2
    Author:     Brian Huffman, TU Muenchen
51956
a4d81cdebf8b better support for domains in Lifting/Transfer = replace Domainp T by the actual invariant in a transferred goal
kuncar
parents: 51955
diff changeset
     3
    Author:     Ondrej Kuncar, TU Muenchen
47325
ec6187036495 new transfer proof method
huffman
parents:
diff changeset
     4
ec6187036495 new transfer proof method
huffman
parents:
diff changeset
     5
Generic theorem transfer method.
ec6187036495 new transfer proof method
huffman
parents:
diff changeset
     6
*)
ec6187036495 new transfer proof method
huffman
parents:
diff changeset
     7
ec6187036495 new transfer proof method
huffman
parents:
diff changeset
     8
signature TRANSFER =
ec6187036495 new transfer proof method
huffman
parents:
diff changeset
     9
sig
56524
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents: 56520
diff changeset
    10
  type pred_data
60220
530112e8c6ba tuned; store pred_simps
kuncar
parents: 60216
diff changeset
    11
  val mk_pred_data: thm -> thm -> thm list -> pred_data
56524
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents: 56520
diff changeset
    12
  val rel_eq_onp: pred_data -> thm
60216
ef4f0b5b925e added pred_def, rel_eq_onp tuned
kuncar
parents: 59642
diff changeset
    13
  val pred_def: pred_data -> thm
60220
530112e8c6ba tuned; store pred_simps
kuncar
parents: 60216
diff changeset
    14
  val pred_simps: pred_data -> thm list
530112e8c6ba tuned; store pred_simps
kuncar
parents: 60216
diff changeset
    15
  val update_pred_simps: thm list -> pred_data -> pred_data
56524
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents: 56520
diff changeset
    16
52883
0a7c97c76f46 expand equalities in the transfer relation in transfer_prover if the relation doesn't follow the functional structure
kuncar
parents: 52358
diff changeset
    17
  val bottom_rewr_conv: thm list -> conv
0a7c97c76f46 expand equalities in the transfer relation in transfer_prover if the relation doesn't follow the functional structure
kuncar
parents: 52358
diff changeset
    18
  val top_rewr_conv: thm list -> conv
60216
ef4f0b5b925e added pred_def, rel_eq_onp tuned
kuncar
parents: 59642
diff changeset
    19
  val top_sweep_rewr_conv: thm list -> conv
52883
0a7c97c76f46 expand equalities in the transfer relation in transfer_prover if the relation doesn't follow the functional structure
kuncar
parents: 52358
diff changeset
    20
47325
ec6187036495 new transfer proof method
huffman
parents:
diff changeset
    21
  val prep_conv: conv
64434
af5235830c16 always expand equalities in the transfer relation in transfer_prover (cf. 0a7c97c76f46)
kuncar
parents: 62958
diff changeset
    22
  val fold_relator_eqs_conv: Proof.context -> conv
af5235830c16 always expand equalities in the transfer relation in transfer_prover (cf. 0a7c97c76f46)
kuncar
parents: 62958
diff changeset
    23
  val unfold_relator_eqs_conv: Proof.context -> conv
51956
a4d81cdebf8b better support for domains in Lifting/Transfer = replace Domainp T by the actual invariant in a transferred goal
kuncar
parents: 51955
diff changeset
    24
  val get_transfer_raw: Proof.context -> thm list
55563
a64d49f49ca3 implement the reflexivity prover as a monotonicity prover that proves R >= op=; derive "reflexivity" rules for relators from mono rules and eq rules
kuncar
parents: 54883
diff changeset
    25
  val get_relator_eq_item_net: Proof.context -> thm Item_Net.T
47503
cb44d09d9d22 add theory data for relator identity rules;
huffman
parents: 47356
diff changeset
    26
  val get_relator_eq: Proof.context -> thm list
49625
06cf80661e7a new get function for non-symmetric relator_eq & tuned
kuncar
parents: 48066
diff changeset
    27
  val get_sym_relator_eq: Proof.context -> thm list
51954
2e3f9e72b8c4 publish a private function
kuncar
parents: 51437
diff changeset
    28
  val get_relator_eq_raw: Proof.context -> thm list
51956
a4d81cdebf8b better support for domains in Lifting/Transfer = replace Domainp T by the actual invariant in a transferred goal
kuncar
parents: 51955
diff changeset
    29
  val get_relator_domain: Proof.context -> thm list
56524
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents: 56520
diff changeset
    30
  val morph_pred_data: morphism -> pred_data -> pred_data
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents: 56520
diff changeset
    31
  val lookup_pred_data: Proof.context -> string -> pred_data option
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents: 56520
diff changeset
    32
  val update_pred_data: string -> pred_data -> Context.generic -> Context.generic
53145
2fb458aceb78 delete corresponding compound lhs and rhs when a transfer rule is deleted; tuned
kuncar
parents: 53144
diff changeset
    33
  val get_compound_lhs: Proof.context -> (term * thm) Item_Net.T
2fb458aceb78 delete corresponding compound lhs and rhs when a transfer rule is deleted; tuned
kuncar
parents: 53144
diff changeset
    34
  val get_compound_rhs: Proof.context -> (term * thm) Item_Net.T
47325
ec6187036495 new transfer proof method
huffman
parents:
diff changeset
    35
  val transfer_add: attribute
ec6187036495 new transfer proof method
huffman
parents:
diff changeset
    36
  val transfer_del: attribute
53649
96814d676c49 public access to the raw transfer rules - for restoring transferring
kuncar
parents: 53145
diff changeset
    37
  val transfer_raw_add: thm -> Context.generic -> Context.generic
96814d676c49 public access to the raw transfer rules - for restoring transferring
kuncar
parents: 53145
diff changeset
    38
  val transfer_raw_del: thm -> Context.generic -> Context.generic
52354
acb4f932dd24 implement 'transferred' attribute for transfer package, with support for monotonicity of !!/==>
huffman
parents: 51996
diff changeset
    39
  val transferred_attribute: thm list -> attribute
52358
f4c4bcb0d564 implement 'untransferred' attribute, which is like 'transferred' but works in the opposite direction
huffman
parents: 52354
diff changeset
    40
  val untransferred_attribute: thm list -> attribute
55731
66df76dd2640 rewrite composition of quotients to a more readable form in a respectfulness goal that is presented to a user
kuncar
parents: 55563
diff changeset
    41
  val prep_transfer_domain_thm: Proof.context -> thm -> thm
51956
a4d81cdebf8b better support for domains in Lifting/Transfer = replace Domainp T by the actual invariant in a transferred goal
kuncar
parents: 51955
diff changeset
    42
  val transfer_domain_add: attribute
a4d81cdebf8b better support for domains in Lifting/Transfer = replace Domainp T by the actual invariant in a transferred goal
kuncar
parents: 51955
diff changeset
    43
  val transfer_domain_del: attribute
52354
acb4f932dd24 implement 'transferred' attribute for transfer package, with support for monotonicity of !!/==>
huffman
parents: 51996
diff changeset
    44
  val transfer_rule_of_term: Proof.context -> bool -> term -> thm
acb4f932dd24 implement 'transferred' attribute for transfer package, with support for monotonicity of !!/==>
huffman
parents: 51996
diff changeset
    45
  val transfer_rule_of_lhs: Proof.context -> term -> thm
55731
66df76dd2640 rewrite composition of quotients to a more readable form in a respectfulness goal that is presented to a user
kuncar
parents: 55563
diff changeset
    46
  val eq_tac: Proof.context -> int -> tactic
61367
46af4f577c7e new methods for debugging transfer and transfer_prover
kuncar
parents: 61366
diff changeset
    47
  val transfer_start_tac: bool -> Proof.context -> int -> tactic
46af4f577c7e new methods for debugging transfer and transfer_prover
kuncar
parents: 61366
diff changeset
    48
  val transfer_prover_start_tac: Proof.context -> int -> tactic
55731
66df76dd2640 rewrite composition of quotients to a more readable form in a respectfulness goal that is presented to a user
kuncar
parents: 55563
diff changeset
    49
  val transfer_step_tac: Proof.context -> int -> tactic
61367
46af4f577c7e new methods for debugging transfer and transfer_prover
kuncar
parents: 61366
diff changeset
    50
  val transfer_end_tac: Proof.context -> int -> tactic
46af4f577c7e new methods for debugging transfer and transfer_prover
kuncar
parents: 61366
diff changeset
    51
  val transfer_prover_end_tac: Proof.context -> int -> tactic
47658
7631f6f7873d enable variant of transfer method that proves an implication instead of an equivalence
huffman
parents: 47635
diff changeset
    52
  val transfer_tac: bool -> Proof.context -> int -> tactic
47635
ebb79474262c rename 'correspondence' method to 'transfer_prover'
huffman
parents: 47618
diff changeset
    53
  val transfer_prover_tac: Proof.context -> int -> tactic
52354
acb4f932dd24 implement 'transferred' attribute for transfer package, with support for monotonicity of !!/==>
huffman
parents: 51996
diff changeset
    54
  val gen_frees_tac: (string * typ) list -> Proof.context -> int -> tactic
47325
ec6187036495 new transfer proof method
huffman
parents:
diff changeset
    55
end
ec6187036495 new transfer proof method
huffman
parents:
diff changeset
    56
ec6187036495 new transfer proof method
huffman
parents:
diff changeset
    57
structure Transfer : TRANSFER =
ec6187036495 new transfer proof method
huffman
parents:
diff changeset
    58
struct
ec6187036495 new transfer proof method
huffman
parents:
diff changeset
    59
69593
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 67664
diff changeset
    60
fun bottom_rewr_conv rewrs = Conv.bottom_conv (K (Conv.try_conv (Conv.rewrs_conv rewrs))) \<^context>
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 67664
diff changeset
    61
fun top_rewr_conv rewrs = Conv.top_conv (K (Conv.try_conv (Conv.rewrs_conv rewrs))) \<^context>
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 67664
diff changeset
    62
fun top_sweep_rewr_conv rewrs = Conv.top_sweep_conv (K (Conv.rewrs_conv rewrs)) \<^context>
60216
ef4f0b5b925e added pred_def, rel_eq_onp tuned
kuncar
parents: 59642
diff changeset
    63
48064
7bd9e18ce058 unify theory-data structures for transfer package
huffman
parents: 47803
diff changeset
    64
(** Theory Data **)
7bd9e18ce058 unify theory-data structures for transfer package
huffman
parents: 47803
diff changeset
    65
59058
a78612c67ec0 renamed "pairself" to "apply2", in accordance to @{apply 2};
wenzelm
parents: 58821
diff changeset
    66
val compound_xhs_empty_net = Item_Net.init (Thm.eq_thm_prop o apply2 snd) (single o fst);
58821
11e226e8a095 modernized setup;
wenzelm
parents: 56722
diff changeset
    67
val rewr_rules = Item_Net.init Thm.eq_thm_prop (single o fst o HOLogic.dest_eq
55563
a64d49f49ca3 implement the reflexivity prover as a monotonicity prover that proves R >= op=; derive "reflexivity" rules for relators from mono rules and eq rules
kuncar
parents: 54883
diff changeset
    68
  o HOLogic.dest_Trueprop o Thm.concl_of);
53145
2fb458aceb78 delete corresponding compound lhs and rhs when a transfer rule is deleted; tuned
kuncar
parents: 53144
diff changeset
    69
60220
530112e8c6ba tuned; store pred_simps
kuncar
parents: 60216
diff changeset
    70
datatype pred_data = PRED_DATA of {pred_def:thm, rel_eq_onp: thm, pred_simps: thm list}
530112e8c6ba tuned; store pred_simps
kuncar
parents: 60216
diff changeset
    71
530112e8c6ba tuned; store pred_simps
kuncar
parents: 60216
diff changeset
    72
fun mk_pred_data pred_def rel_eq_onp pred_simps = PRED_DATA {pred_def = pred_def, 
530112e8c6ba tuned; store pred_simps
kuncar
parents: 60216
diff changeset
    73
  rel_eq_onp = rel_eq_onp, pred_simps = pred_simps}
530112e8c6ba tuned; store pred_simps
kuncar
parents: 60216
diff changeset
    74
530112e8c6ba tuned; store pred_simps
kuncar
parents: 60216
diff changeset
    75
fun map_pred_data' f1 f2 f3 (PRED_DATA {pred_def, rel_eq_onp, pred_simps}) =
530112e8c6ba tuned; store pred_simps
kuncar
parents: 60216
diff changeset
    76
  PRED_DATA {pred_def = f1 pred_def, rel_eq_onp = f2 rel_eq_onp, pred_simps = f3 pred_simps}
56524
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents: 56520
diff changeset
    77
60220
530112e8c6ba tuned; store pred_simps
kuncar
parents: 60216
diff changeset
    78
fun rep_pred_data (PRED_DATA p) = p
530112e8c6ba tuned; store pred_simps
kuncar
parents: 60216
diff changeset
    79
val rel_eq_onp = #rel_eq_onp o rep_pred_data
530112e8c6ba tuned; store pred_simps
kuncar
parents: 60216
diff changeset
    80
val pred_def = #pred_def o rep_pred_data
530112e8c6ba tuned; store pred_simps
kuncar
parents: 60216
diff changeset
    81
val pred_simps = #pred_simps o rep_pred_data
530112e8c6ba tuned; store pred_simps
kuncar
parents: 60216
diff changeset
    82
fun update_pred_simps new_pred_data = map_pred_data' I I (K new_pred_data)
530112e8c6ba tuned; store pred_simps
kuncar
parents: 60216
diff changeset
    83
56524
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents: 56520
diff changeset
    84
48064
7bd9e18ce058 unify theory-data structures for transfer package
huffman
parents: 47803
diff changeset
    85
structure Data = Generic_Data
47325
ec6187036495 new transfer proof method
huffman
parents:
diff changeset
    86
(
48064
7bd9e18ce058 unify theory-data structures for transfer package
huffman
parents: 47803
diff changeset
    87
  type T =
7bd9e18ce058 unify theory-data structures for transfer package
huffman
parents: 47803
diff changeset
    88
    { transfer_raw : thm Item_Net.T,
48065
8aa05d38299a transfer method now avoids generalizing over free variables that are known to appear in registered transfer rules
huffman
parents: 48064
diff changeset
    89
      known_frees : (string * typ) list,
53145
2fb458aceb78 delete corresponding compound lhs and rhs when a transfer rule is deleted; tuned
kuncar
parents: 53144
diff changeset
    90
      compound_lhs : (term * thm) Item_Net.T,
2fb458aceb78 delete corresponding compound lhs and rhs when a transfer rule is deleted; tuned
kuncar
parents: 53144
diff changeset
    91
      compound_rhs : (term * thm) Item_Net.T,
51437
8739f8abbecb fixing transfer tactic - unfold fully identity relation by using relator_eq
kuncar
parents: 51374
diff changeset
    92
      relator_eq : thm Item_Net.T,
51956
a4d81cdebf8b better support for domains in Lifting/Transfer = replace Domainp T by the actual invariant in a transferred goal
kuncar
parents: 51955
diff changeset
    93
      relator_eq_raw : thm Item_Net.T,
56524
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents: 56520
diff changeset
    94
      relator_domain : thm Item_Net.T,
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents: 56520
diff changeset
    95
      pred_data : pred_data Symtab.table }
48064
7bd9e18ce058 unify theory-data structures for transfer package
huffman
parents: 47803
diff changeset
    96
  val empty =
52354
acb4f932dd24 implement 'transferred' attribute for transfer package, with support for monotonicity of !!/==>
huffman
parents: 51996
diff changeset
    97
    { transfer_raw = Thm.intro_rules,
48065
8aa05d38299a transfer method now avoids generalizing over free variables that are known to appear in registered transfer rules
huffman
parents: 48064
diff changeset
    98
      known_frees = [],
53145
2fb458aceb78 delete corresponding compound lhs and rhs when a transfer rule is deleted; tuned
kuncar
parents: 53144
diff changeset
    99
      compound_lhs = compound_xhs_empty_net,
2fb458aceb78 delete corresponding compound lhs and rhs when a transfer rule is deleted; tuned
kuncar
parents: 53144
diff changeset
   100
      compound_rhs = compound_xhs_empty_net,
55563
a64d49f49ca3 implement the reflexivity prover as a monotonicity prover that proves R >= op=; derive "reflexivity" rules for relators from mono rules and eq rules
kuncar
parents: 54883
diff changeset
   101
      relator_eq = rewr_rules,
51956
a4d81cdebf8b better support for domains in Lifting/Transfer = replace Domainp T by the actual invariant in a transferred goal
kuncar
parents: 51955
diff changeset
   102
      relator_eq_raw = Thm.full_rules,
56524
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents: 56520
diff changeset
   103
      relator_domain = Thm.full_rules,
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents: 56520
diff changeset
   104
      pred_data = Symtab.empty }
48064
7bd9e18ce058 unify theory-data structures for transfer package
huffman
parents: 47803
diff changeset
   105
  val extend = I
48066
c6783c9b87bf transfer method now handles transfer rules for compound terms, e.g. locale-defined constants with hidden parameters
huffman
parents: 48065
diff changeset
   106
  fun merge
c6783c9b87bf transfer method now handles transfer rules for compound terms, e.g. locale-defined constants with hidden parameters
huffman
parents: 48065
diff changeset
   107
    ( { transfer_raw = t1, known_frees = k1,
52354
acb4f932dd24 implement 'transferred' attribute for transfer package, with support for monotonicity of !!/==>
huffman
parents: 51996
diff changeset
   108
        compound_lhs = l1,
51437
8739f8abbecb fixing transfer tactic - unfold fully identity relation by using relator_eq
kuncar
parents: 51374
diff changeset
   109
        compound_rhs = c1, relator_eq = r1,
56524
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents: 56520
diff changeset
   110
        relator_eq_raw = rw1, relator_domain = rd1,
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents: 56520
diff changeset
   111
        pred_data = pd1 },
48066
c6783c9b87bf transfer method now handles transfer rules for compound terms, e.g. locale-defined constants with hidden parameters
huffman
parents: 48065
diff changeset
   112
      { transfer_raw = t2, known_frees = k2,
52354
acb4f932dd24 implement 'transferred' attribute for transfer package, with support for monotonicity of !!/==>
huffman
parents: 51996
diff changeset
   113
        compound_lhs = l2,
51437
8739f8abbecb fixing transfer tactic - unfold fully identity relation by using relator_eq
kuncar
parents: 51374
diff changeset
   114
        compound_rhs = c2, relator_eq = r2,
56524
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents: 56520
diff changeset
   115
        relator_eq_raw = rw2, relator_domain = rd2,
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents: 56520
diff changeset
   116
        pred_data = pd2 } ) =
48064
7bd9e18ce058 unify theory-data structures for transfer package
huffman
parents: 47803
diff changeset
   117
    { transfer_raw = Item_Net.merge (t1, t2),
48065
8aa05d38299a transfer method now avoids generalizing over free variables that are known to appear in registered transfer rules
huffman
parents: 48064
diff changeset
   118
      known_frees = Library.merge (op =) (k1, k2),
53145
2fb458aceb78 delete corresponding compound lhs and rhs when a transfer rule is deleted; tuned
kuncar
parents: 53144
diff changeset
   119
      compound_lhs = Item_Net.merge (l1, l2),
2fb458aceb78 delete corresponding compound lhs and rhs when a transfer rule is deleted; tuned
kuncar
parents: 53144
diff changeset
   120
      compound_rhs = Item_Net.merge (c1, c2),
51437
8739f8abbecb fixing transfer tactic - unfold fully identity relation by using relator_eq
kuncar
parents: 51374
diff changeset
   121
      relator_eq = Item_Net.merge (r1, r2),
51956
a4d81cdebf8b better support for domains in Lifting/Transfer = replace Domainp T by the actual invariant in a transferred goal
kuncar
parents: 51955
diff changeset
   122
      relator_eq_raw = Item_Net.merge (rw1, rw2),
56524
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents: 56520
diff changeset
   123
      relator_domain = Item_Net.merge (rd1, rd2),
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents: 56520
diff changeset
   124
      pred_data = Symtab.merge (K true) (pd1, pd2) }
47325
ec6187036495 new transfer proof method
huffman
parents:
diff changeset
   125
)
ec6187036495 new transfer proof method
huffman
parents:
diff changeset
   126
70316
c61b7af108a6 misc tuning and clarification, notably wrt. flow of context;
wenzelm
parents: 70159
diff changeset
   127
val get_transfer_raw = Item_Net.content o #transfer_raw o Data.get o Context.Proof
51956
a4d81cdebf8b better support for domains in Lifting/Transfer = replace Domainp T by the actual invariant in a transferred goal
kuncar
parents: 51955
diff changeset
   128
70316
c61b7af108a6 misc tuning and clarification, notably wrt. flow of context;
wenzelm
parents: 70159
diff changeset
   129
val get_known_frees = #known_frees o Data.get o Context.Proof
51956
a4d81cdebf8b better support for domains in Lifting/Transfer = replace Domainp T by the actual invariant in a transferred goal
kuncar
parents: 51955
diff changeset
   130
70316
c61b7af108a6 misc tuning and clarification, notably wrt. flow of context;
wenzelm
parents: 70159
diff changeset
   131
val get_compound_lhs = #compound_lhs o Data.get o Context.Proof
52354
acb4f932dd24 implement 'transferred' attribute for transfer package, with support for monotonicity of !!/==>
huffman
parents: 51996
diff changeset
   132
70316
c61b7af108a6 misc tuning and clarification, notably wrt. flow of context;
wenzelm
parents: 70159
diff changeset
   133
val get_compound_rhs = #compound_rhs o Data.get o Context.Proof
51956
a4d81cdebf8b better support for domains in Lifting/Transfer = replace Domainp T by the actual invariant in a transferred goal
kuncar
parents: 51955
diff changeset
   134
70316
c61b7af108a6 misc tuning and clarification, notably wrt. flow of context;
wenzelm
parents: 70159
diff changeset
   135
val get_relator_eq_item_net = #relator_eq o Data.get o Context.Proof
55563
a64d49f49ca3 implement the reflexivity prover as a monotonicity prover that proves R >= op=; derive "reflexivity" rules for relators from mono rules and eq rules
kuncar
parents: 54883
diff changeset
   136
70316
c61b7af108a6 misc tuning and clarification, notably wrt. flow of context;
wenzelm
parents: 70159
diff changeset
   137
val get_relator_eq =
c61b7af108a6 misc tuning and clarification, notably wrt. flow of context;
wenzelm
parents: 70159
diff changeset
   138
  map safe_mk_meta_eq o Item_Net.content o #relator_eq o Data.get o Context.Proof
49625
06cf80661e7a new get function for non-symmetric relator_eq & tuned
kuncar
parents: 48066
diff changeset
   139
70316
c61b7af108a6 misc tuning and clarification, notably wrt. flow of context;
wenzelm
parents: 70159
diff changeset
   140
val get_sym_relator_eq =
c61b7af108a6 misc tuning and clarification, notably wrt. flow of context;
wenzelm
parents: 70159
diff changeset
   141
  map (Thm.symmetric o safe_mk_meta_eq) o Item_Net.content o #relator_eq o Data.get o Context.Proof
48064
7bd9e18ce058 unify theory-data structures for transfer package
huffman
parents: 47803
diff changeset
   142
70316
c61b7af108a6 misc tuning and clarification, notably wrt. flow of context;
wenzelm
parents: 70159
diff changeset
   143
val get_relator_eq_raw = Item_Net.content o #relator_eq_raw o Data.get o Context.Proof
51437
8739f8abbecb fixing transfer tactic - unfold fully identity relation by using relator_eq
kuncar
parents: 51374
diff changeset
   144
70316
c61b7af108a6 misc tuning and clarification, notably wrt. flow of context;
wenzelm
parents: 70159
diff changeset
   145
val get_relator_domain = Item_Net.content o #relator_domain o Data.get o Context.Proof
48065
8aa05d38299a transfer method now avoids generalizing over free variables that are known to appear in registered transfer rules
huffman
parents: 48064
diff changeset
   146
70316
c61b7af108a6 misc tuning and clarification, notably wrt. flow of context;
wenzelm
parents: 70159
diff changeset
   147
val get_pred_data = #pred_data o Data.get o Context.Proof
56524
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents: 56520
diff changeset
   148
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents: 56520
diff changeset
   149
fun map_data f1 f2 f3 f4 f5 f6 f7 f8
52354
acb4f932dd24 implement 'transferred' attribute for transfer package, with support for monotonicity of !!/==>
huffman
parents: 51996
diff changeset
   150
  { transfer_raw, known_frees, compound_lhs, compound_rhs,
56524
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents: 56520
diff changeset
   151
    relator_eq, relator_eq_raw, relator_domain, pred_data } =
48065
8aa05d38299a transfer method now avoids generalizing over free variables that are known to appear in registered transfer rules
huffman
parents: 48064
diff changeset
   152
  { transfer_raw = f1 transfer_raw,
8aa05d38299a transfer method now avoids generalizing over free variables that are known to appear in registered transfer rules
huffman
parents: 48064
diff changeset
   153
    known_frees = f2 known_frees,
52354
acb4f932dd24 implement 'transferred' attribute for transfer package, with support for monotonicity of !!/==>
huffman
parents: 51996
diff changeset
   154
    compound_lhs = f3 compound_lhs,
acb4f932dd24 implement 'transferred' attribute for transfer package, with support for monotonicity of !!/==>
huffman
parents: 51996
diff changeset
   155
    compound_rhs = f4 compound_rhs,
acb4f932dd24 implement 'transferred' attribute for transfer package, with support for monotonicity of !!/==>
huffman
parents: 51996
diff changeset
   156
    relator_eq = f5 relator_eq,
acb4f932dd24 implement 'transferred' attribute for transfer package, with support for monotonicity of !!/==>
huffman
parents: 51996
diff changeset
   157
    relator_eq_raw = f6 relator_eq_raw,
56524
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents: 56520
diff changeset
   158
    relator_domain = f7 relator_domain,
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents: 56520
diff changeset
   159
    pred_data = f8 pred_data }
48064
7bd9e18ce058 unify theory-data structures for transfer package
huffman
parents: 47803
diff changeset
   160
56524
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents: 56520
diff changeset
   161
fun map_transfer_raw   f = map_data f I I I I I I I
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents: 56520
diff changeset
   162
fun map_known_frees    f = map_data I f I I I I I I
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents: 56520
diff changeset
   163
fun map_compound_lhs   f = map_data I I f I I I I I
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents: 56520
diff changeset
   164
fun map_compound_rhs   f = map_data I I I f I I I I
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents: 56520
diff changeset
   165
fun map_relator_eq     f = map_data I I I I f I I I
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents: 56520
diff changeset
   166
fun map_relator_eq_raw f = map_data I I I I I f I I
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents: 56520
diff changeset
   167
fun map_relator_domain f = map_data I I I I I I f I
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents: 56520
diff changeset
   168
fun map_pred_data      f = map_data I I I I I I I f
47503
cb44d09d9d22 add theory data for relator identity rules;
huffman
parents: 47356
diff changeset
   169
48065
8aa05d38299a transfer method now avoids generalizing over free variables that are known to appear in registered transfer rules
huffman
parents: 48064
diff changeset
   170
fun add_transfer_thm thm = Data.map
8aa05d38299a transfer method now avoids generalizing over free variables that are known to appear in registered transfer rules
huffman
parents: 48064
diff changeset
   171
  (map_transfer_raw (Item_Net.update thm) o
52354
acb4f932dd24 implement 'transferred' attribute for transfer package, with support for monotonicity of !!/==>
huffman
parents: 51996
diff changeset
   172
   map_compound_lhs
acb4f932dd24 implement 'transferred' attribute for transfer package, with support for monotonicity of !!/==>
huffman
parents: 51996
diff changeset
   173
     (case HOLogic.dest_Trueprop (Thm.concl_of thm) of
69593
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 67664
diff changeset
   174
        Const (\<^const_name>\<open>Rel\<close>, _) $ _ $ (lhs as (_ $ _)) $ _ =>
53145
2fb458aceb78 delete corresponding compound lhs and rhs when a transfer rule is deleted; tuned
kuncar
parents: 53144
diff changeset
   175
          Item_Net.update (lhs, thm)
52354
acb4f932dd24 implement 'transferred' attribute for transfer package, with support for monotonicity of !!/==>
huffman
parents: 51996
diff changeset
   176
      | _ => I) o
48066
c6783c9b87bf transfer method now handles transfer rules for compound terms, e.g. locale-defined constants with hidden parameters
huffman
parents: 48065
diff changeset
   177
   map_compound_rhs
c6783c9b87bf transfer method now handles transfer rules for compound terms, e.g. locale-defined constants with hidden parameters
huffman
parents: 48065
diff changeset
   178
     (case HOLogic.dest_Trueprop (Thm.concl_of thm) of
69593
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 67664
diff changeset
   179
        Const (\<^const_name>\<open>Rel\<close>, _) $ _ $ _ $ (rhs as (_ $ _)) =>
53145
2fb458aceb78 delete corresponding compound lhs and rhs when a transfer rule is deleted; tuned
kuncar
parents: 53144
diff changeset
   180
          Item_Net.update (rhs, thm)
48066
c6783c9b87bf transfer method now handles transfer rules for compound terms, e.g. locale-defined constants with hidden parameters
huffman
parents: 48065
diff changeset
   181
      | _ => I) o
48065
8aa05d38299a transfer method now avoids generalizing over free variables that are known to appear in registered transfer rules
huffman
parents: 48064
diff changeset
   182
   map_known_frees (Term.add_frees (Thm.concl_of thm)))
8aa05d38299a transfer method now avoids generalizing over free variables that are known to appear in registered transfer rules
huffman
parents: 48064
diff changeset
   183
58821
11e226e8a095 modernized setup;
wenzelm
parents: 56722
diff changeset
   184
fun del_transfer_thm thm = Data.map
53145
2fb458aceb78 delete corresponding compound lhs and rhs when a transfer rule is deleted; tuned
kuncar
parents: 53144
diff changeset
   185
  (map_transfer_raw (Item_Net.remove thm) o
2fb458aceb78 delete corresponding compound lhs and rhs when a transfer rule is deleted; tuned
kuncar
parents: 53144
diff changeset
   186
   map_compound_lhs
2fb458aceb78 delete corresponding compound lhs and rhs when a transfer rule is deleted; tuned
kuncar
parents: 53144
diff changeset
   187
     (case HOLogic.dest_Trueprop (Thm.concl_of thm) of
69593
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 67664
diff changeset
   188
        Const (\<^const_name>\<open>Rel\<close>, _) $ _ $ (lhs as (_ $ _)) $ _ =>
53145
2fb458aceb78 delete corresponding compound lhs and rhs when a transfer rule is deleted; tuned
kuncar
parents: 53144
diff changeset
   189
          Item_Net.remove (lhs, thm)
2fb458aceb78 delete corresponding compound lhs and rhs when a transfer rule is deleted; tuned
kuncar
parents: 53144
diff changeset
   190
      | _ => I) o
2fb458aceb78 delete corresponding compound lhs and rhs when a transfer rule is deleted; tuned
kuncar
parents: 53144
diff changeset
   191
   map_compound_rhs
2fb458aceb78 delete corresponding compound lhs and rhs when a transfer rule is deleted; tuned
kuncar
parents: 53144
diff changeset
   192
     (case HOLogic.dest_Trueprop (Thm.concl_of thm) of
69593
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 67664
diff changeset
   193
        Const (\<^const_name>\<open>Rel\<close>, _) $ _ $ _ $ (rhs as (_ $ _)) =>
53145
2fb458aceb78 delete corresponding compound lhs and rhs when a transfer rule is deleted; tuned
kuncar
parents: 53144
diff changeset
   194
          Item_Net.remove (rhs, thm)
2fb458aceb78 delete corresponding compound lhs and rhs when a transfer rule is deleted; tuned
kuncar
parents: 53144
diff changeset
   195
      | _ => I))
48064
7bd9e18ce058 unify theory-data structures for transfer package
huffman
parents: 47803
diff changeset
   196
53649
96814d676c49 public access to the raw transfer rules - for restoring transferring
kuncar
parents: 53145
diff changeset
   197
fun transfer_raw_add thm ctxt = add_transfer_thm thm ctxt
96814d676c49 public access to the raw transfer rules - for restoring transferring
kuncar
parents: 53145
diff changeset
   198
fun transfer_raw_del thm ctxt = del_transfer_thm thm ctxt
96814d676c49 public access to the raw transfer rules - for restoring transferring
kuncar
parents: 53145
diff changeset
   199
47325
ec6187036495 new transfer proof method
huffman
parents:
diff changeset
   200
(** Conversions **)
ec6187036495 new transfer proof method
huffman
parents:
diff changeset
   201
58821
11e226e8a095 modernized setup;
wenzelm
parents: 56722
diff changeset
   202
fun transfer_rel_conv conv =
52883
0a7c97c76f46 expand equalities in the transfer relation in transfer_prover if the relation doesn't follow the functional structure
kuncar
parents: 52358
diff changeset
   203
  Conv.concl_conv ~1 (HOLogic.Trueprop_conv (Conv.fun2_conv (Conv.arg_conv conv)))
0a7c97c76f46 expand equalities in the transfer relation in transfer_prover if the relation doesn't follow the functional structure
kuncar
parents: 52358
diff changeset
   204
47325
ec6187036495 new transfer proof method
huffman
parents:
diff changeset
   205
val Rel_rule = Thm.symmetric @{thm Rel_def}
ec6187036495 new transfer proof method
huffman
parents:
diff changeset
   206
ec6187036495 new transfer proof method
huffman
parents:
diff changeset
   207
fun Rel_conv ct =
70159
57503fe1b0ff tuned signature;
wenzelm
parents: 70154
diff changeset
   208
  let val (cT, cT') = Thm.dest_funT (Thm.ctyp_of_cterm ct)
57503fe1b0ff tuned signature;
wenzelm
parents: 70154
diff changeset
   209
      val (cU, _) = Thm.dest_funT cT'
60801
7664e0916eec tuned signature;
wenzelm
parents: 60752
diff changeset
   210
  in Thm.instantiate' [SOME cT, SOME cU] [SOME ct] Rel_rule end
47325
ec6187036495 new transfer proof method
huffman
parents:
diff changeset
   211
47635
ebb79474262c rename 'correspondence' method to 'transfer_prover'
huffman
parents: 47618
diff changeset
   212
(* Conversion to preprocess a transfer rule *)
51955
04d9381bebff try to detect assumptions of transfer rules that are in a shape of a transfer rule
kuncar
parents: 51954
diff changeset
   213
fun safe_Rel_conv ct =
51956
a4d81cdebf8b better support for domains in Lifting/Transfer = replace Domainp T by the actual invariant in a transferred goal
kuncar
parents: 51955
diff changeset
   214
  Conv.try_conv (HOLogic.Trueprop_conv (Conv.fun_conv (Conv.fun_conv Rel_conv))) ct
51955
04d9381bebff try to detect assumptions of transfer rules that are in a shape of a transfer rule
kuncar
parents: 51954
diff changeset
   215
47325
ec6187036495 new transfer proof method
huffman
parents:
diff changeset
   216
fun prep_conv ct = (
51955
04d9381bebff try to detect assumptions of transfer rules that are in a shape of a transfer rule
kuncar
parents: 51954
diff changeset
   217
      Conv.implies_conv safe_Rel_conv prep_conv
47325
ec6187036495 new transfer proof method
huffman
parents:
diff changeset
   218
      else_conv
51955
04d9381bebff try to detect assumptions of transfer rules that are in a shape of a transfer rule
kuncar
parents: 51954
diff changeset
   219
      safe_Rel_conv
47325
ec6187036495 new transfer proof method
huffman
parents:
diff changeset
   220
      else_conv
ec6187036495 new transfer proof method
huffman
parents:
diff changeset
   221
      Conv.all_conv) ct
ec6187036495 new transfer proof method
huffman
parents:
diff changeset
   222
64434
af5235830c16 always expand equalities in the transfer relation in transfer_prover (cf. 0a7c97c76f46)
kuncar
parents: 62958
diff changeset
   223
fun fold_relator_eqs_conv ctxt ct = (bottom_rewr_conv (get_relator_eq ctxt)) ct;
af5235830c16 always expand equalities in the transfer relation in transfer_prover (cf. 0a7c97c76f46)
kuncar
parents: 62958
diff changeset
   224
fun unfold_relator_eqs_conv ctxt ct = (top_rewr_conv (get_sym_relator_eq ctxt)) ct;
af5235830c16 always expand equalities in the transfer relation in transfer_prover (cf. 0a7c97c76f46)
kuncar
parents: 62958
diff changeset
   225
af5235830c16 always expand equalities in the transfer relation in transfer_prover (cf. 0a7c97c76f46)
kuncar
parents: 62958
diff changeset
   226
49975
faf4afed009f transfer package: more flexible handling of equality relations using is_equality predicate
huffman
parents: 49625
diff changeset
   227
(** Replacing explicit equalities with is_equality premises **)
faf4afed009f transfer package: more flexible handling of equality relations using is_equality predicate
huffman
parents: 49625
diff changeset
   228
faf4afed009f transfer package: more flexible handling of equality relations using is_equality predicate
huffman
parents: 49625
diff changeset
   229
fun mk_is_equality t =
69593
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 67664
diff changeset
   230
  Const (\<^const_name>\<open>is_equality\<close>, Term.fastype_of t --> HOLogic.boolT) $ t
47325
ec6187036495 new transfer proof method
huffman
parents:
diff changeset
   231
49975
faf4afed009f transfer package: more flexible handling of equality relations using is_equality predicate
huffman
parents: 49625
diff changeset
   232
val is_equality_lemma =
67399
eab6ce8368fa ran isabelle update_op on all sources
nipkow
parents: 64434
diff changeset
   233
  @{lemma "(!!R. is_equality R ==> PROP (P R)) == PROP (P (=))"
49975
faf4afed009f transfer package: more flexible handling of equality relations using is_equality predicate
huffman
parents: 49625
diff changeset
   234
    by (unfold is_equality_def, rule, drule meta_spec,
faf4afed009f transfer package: more flexible handling of equality relations using is_equality predicate
huffman
parents: 49625
diff changeset
   235
      erule meta_mp, rule refl, simp)}
47325
ec6187036495 new transfer proof method
huffman
parents:
diff changeset
   236
54742
7a86358a3c0b proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
wenzelm
parents: 53649
diff changeset
   237
fun gen_abstract_equalities ctxt (dest : term -> term * (term -> term)) thm =
47325
ec6187036495 new transfer proof method
huffman
parents:
diff changeset
   238
  let
49975
faf4afed009f transfer package: more flexible handling of equality relations using is_equality predicate
huffman
parents: 49625
diff changeset
   239
    val prop = Thm.prop_of thm
faf4afed009f transfer package: more flexible handling of equality relations using is_equality predicate
huffman
parents: 49625
diff changeset
   240
    val (t, mk_prop') = dest prop
67399
eab6ce8368fa ran isabelle update_op on all sources
nipkow
parents: 64434
diff changeset
   241
    (* Only consider "(=)" at non-base types *)
69593
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 67664
diff changeset
   242
    fun is_eq (Const (\<^const_name>\<open>HOL.eq\<close>, Type ("fun", [T, _]))) =
52354
acb4f932dd24 implement 'transferred' attribute for transfer package, with support for monotonicity of !!/==>
huffman
parents: 51996
diff changeset
   243
        (case T of Type (_, []) => false | _ => true)
acb4f932dd24 implement 'transferred' attribute for transfer package, with support for monotonicity of !!/==>
huffman
parents: 51996
diff changeset
   244
      | is_eq _ = false
acb4f932dd24 implement 'transferred' attribute for transfer package, with support for monotonicity of !!/==>
huffman
parents: 51996
diff changeset
   245
    val add_eqs = Term.fold_aterms (fn t => if is_eq t then insert (op =) t else I)
49975
faf4afed009f transfer package: more flexible handling of equality relations using is_equality predicate
huffman
parents: 49625
diff changeset
   246
    val eq_consts = rev (add_eqs t [])
faf4afed009f transfer package: more flexible handling of equality relations using is_equality predicate
huffman
parents: 49625
diff changeset
   247
    val eqTs = map (snd o dest_Const) eq_consts
faf4afed009f transfer package: more flexible handling of equality relations using is_equality predicate
huffman
parents: 49625
diff changeset
   248
    val used = Term.add_free_names prop []
faf4afed009f transfer package: more flexible handling of equality relations using is_equality predicate
huffman
parents: 49625
diff changeset
   249
    val names = map (K "") eqTs |> Name.variant_list used
faf4afed009f transfer package: more flexible handling of equality relations using is_equality predicate
huffman
parents: 49625
diff changeset
   250
    val frees = map Free (names ~~ eqTs)
faf4afed009f transfer package: more flexible handling of equality relations using is_equality predicate
huffman
parents: 49625
diff changeset
   251
    val prems = map (HOLogic.mk_Trueprop o mk_is_equality) frees
faf4afed009f transfer package: more flexible handling of equality relations using is_equality predicate
huffman
parents: 49625
diff changeset
   252
    val prop1 = mk_prop' (Term.subst_atomic (eq_consts ~~ frees) t)
faf4afed009f transfer package: more flexible handling of equality relations using is_equality predicate
huffman
parents: 49625
diff changeset
   253
    val prop2 = fold Logic.all frees (Logic.list_implies (prems, prop1))
59642
929984c529d3 clarified context;
wenzelm
parents: 59621
diff changeset
   254
    val cprop = Thm.cterm_of ctxt prop2
54742
7a86358a3c0b proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
wenzelm
parents: 53649
diff changeset
   255
    val equal_thm = Raw_Simplifier.rewrite ctxt false [is_equality_lemma] cprop
49975
faf4afed009f transfer package: more flexible handling of equality relations using is_equality predicate
huffman
parents: 49625
diff changeset
   256
    fun forall_elim thm = Thm.forall_elim_vars (Thm.maxidx_of thm + 1) thm
47325
ec6187036495 new transfer proof method
huffman
parents:
diff changeset
   257
  in
49975
faf4afed009f transfer package: more flexible handling of equality relations using is_equality predicate
huffman
parents: 49625
diff changeset
   258
    forall_elim (thm COMP (equal_thm COMP @{thm equal_elim_rule2}))
faf4afed009f transfer package: more flexible handling of equality relations using is_equality predicate
huffman
parents: 49625
diff changeset
   259
  end
faf4afed009f transfer package: more flexible handling of equality relations using is_equality predicate
huffman
parents: 49625
diff changeset
   260
    handle TERM _ => thm
faf4afed009f transfer package: more flexible handling of equality relations using is_equality predicate
huffman
parents: 49625
diff changeset
   261
52884
34c47bc771f2 contract equalities in transfer and transfer domain rules when they are registered
kuncar
parents: 52883
diff changeset
   262
fun abstract_equalities_transfer ctxt thm =
49975
faf4afed009f transfer package: more flexible handling of equality relations using is_equality predicate
huffman
parents: 49625
diff changeset
   263
  let
faf4afed009f transfer package: more flexible handling of equality relations using is_equality predicate
huffman
parents: 49625
diff changeset
   264
    fun dest prop =
faf4afed009f transfer package: more flexible handling of equality relations using is_equality predicate
huffman
parents: 49625
diff changeset
   265
      let
faf4afed009f transfer package: more flexible handling of equality relations using is_equality predicate
huffman
parents: 49625
diff changeset
   266
        val prems = Logic.strip_imp_prems prop
faf4afed009f transfer package: more flexible handling of equality relations using is_equality predicate
huffman
parents: 49625
diff changeset
   267
        val concl = HOLogic.dest_Trueprop (Logic.strip_imp_concl prop)
faf4afed009f transfer package: more flexible handling of equality relations using is_equality predicate
huffman
parents: 49625
diff changeset
   268
        val ((rel, x), y) = apfst Term.dest_comb (Term.dest_comb concl)
faf4afed009f transfer package: more flexible handling of equality relations using is_equality predicate
huffman
parents: 49625
diff changeset
   269
      in
faf4afed009f transfer package: more flexible handling of equality relations using is_equality predicate
huffman
parents: 49625
diff changeset
   270
        (rel, fn rel' =>
faf4afed009f transfer package: more flexible handling of equality relations using is_equality predicate
huffman
parents: 49625
diff changeset
   271
          Logic.list_implies (prems, HOLogic.mk_Trueprop (rel' $ x $ y)))
faf4afed009f transfer package: more flexible handling of equality relations using is_equality predicate
huffman
parents: 49625
diff changeset
   272
      end
58821
11e226e8a095 modernized setup;
wenzelm
parents: 56722
diff changeset
   273
    val contracted_eq_thm =
64434
af5235830c16 always expand equalities in the transfer relation in transfer_prover (cf. 0a7c97c76f46)
kuncar
parents: 62958
diff changeset
   274
      Conv.fconv_rule (transfer_rel_conv (fold_relator_eqs_conv ctxt)) thm
52884
34c47bc771f2 contract equalities in transfer and transfer domain rules when they are registered
kuncar
parents: 52883
diff changeset
   275
      handle CTERM _ => thm
49975
faf4afed009f transfer package: more flexible handling of equality relations using is_equality predicate
huffman
parents: 49625
diff changeset
   276
  in
54742
7a86358a3c0b proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
wenzelm
parents: 53649
diff changeset
   277
    gen_abstract_equalities ctxt dest contracted_eq_thm
49975
faf4afed009f transfer package: more flexible handling of equality relations using is_equality predicate
huffman
parents: 49625
diff changeset
   278
  end
faf4afed009f transfer package: more flexible handling of equality relations using is_equality predicate
huffman
parents: 49625
diff changeset
   279
54742
7a86358a3c0b proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
wenzelm
parents: 53649
diff changeset
   280
fun abstract_equalities_relator_eq ctxt rel_eq_thm =
7a86358a3c0b proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
wenzelm
parents: 53649
diff changeset
   281
  gen_abstract_equalities ctxt (fn x => (x, I))
49975
faf4afed009f transfer package: more flexible handling of equality relations using is_equality predicate
huffman
parents: 49625
diff changeset
   282
    (rel_eq_thm RS @{thm is_equality_def [THEN iffD2]})
faf4afed009f transfer package: more flexible handling of equality relations using is_equality predicate
huffman
parents: 49625
diff changeset
   283
52884
34c47bc771f2 contract equalities in transfer and transfer domain rules when they are registered
kuncar
parents: 52883
diff changeset
   284
fun abstract_equalities_domain ctxt thm =
51956
a4d81cdebf8b better support for domains in Lifting/Transfer = replace Domainp T by the actual invariant in a transferred goal
kuncar
parents: 51955
diff changeset
   285
  let
a4d81cdebf8b better support for domains in Lifting/Transfer = replace Domainp T by the actual invariant in a transferred goal
kuncar
parents: 51955
diff changeset
   286
    fun dest prop =
a4d81cdebf8b better support for domains in Lifting/Transfer = replace Domainp T by the actual invariant in a transferred goal
kuncar
parents: 51955
diff changeset
   287
      let
a4d81cdebf8b better support for domains in Lifting/Transfer = replace Domainp T by the actual invariant in a transferred goal
kuncar
parents: 51955
diff changeset
   288
        val prems = Logic.strip_imp_prems prop
a4d81cdebf8b better support for domains in Lifting/Transfer = replace Domainp T by the actual invariant in a transferred goal
kuncar
parents: 51955
diff changeset
   289
        val concl = HOLogic.dest_Trueprop (Logic.strip_imp_concl prop)
52884
34c47bc771f2 contract equalities in transfer and transfer domain rules when they are registered
kuncar
parents: 52883
diff changeset
   290
        val ((eq, dom), y) = apfst Term.dest_comb (Term.dest_comb concl)
51956
a4d81cdebf8b better support for domains in Lifting/Transfer = replace Domainp T by the actual invariant in a transferred goal
kuncar
parents: 51955
diff changeset
   291
      in
52884
34c47bc771f2 contract equalities in transfer and transfer domain rules when they are registered
kuncar
parents: 52883
diff changeset
   292
        (dom, fn dom' => Logic.list_implies (prems, HOLogic.mk_Trueprop (eq $ dom' $ y)))
51956
a4d81cdebf8b better support for domains in Lifting/Transfer = replace Domainp T by the actual invariant in a transferred goal
kuncar
parents: 51955
diff changeset
   293
      end
58821
11e226e8a095 modernized setup;
wenzelm
parents: 56722
diff changeset
   294
    fun transfer_rel_conv conv =
52884
34c47bc771f2 contract equalities in transfer and transfer domain rules when they are registered
kuncar
parents: 52883
diff changeset
   295
      Conv.concl_conv ~1 (HOLogic.Trueprop_conv (Conv.arg1_conv (Conv.arg_conv conv)))
58821
11e226e8a095 modernized setup;
wenzelm
parents: 56722
diff changeset
   296
    val contracted_eq_thm =
64434
af5235830c16 always expand equalities in the transfer relation in transfer_prover (cf. 0a7c97c76f46)
kuncar
parents: 62958
diff changeset
   297
      Conv.fconv_rule (transfer_rel_conv (fold_relator_eqs_conv ctxt)) thm
51956
a4d81cdebf8b better support for domains in Lifting/Transfer = replace Domainp T by the actual invariant in a transferred goal
kuncar
parents: 51955
diff changeset
   298
  in
54742
7a86358a3c0b proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
wenzelm
parents: 53649
diff changeset
   299
    gen_abstract_equalities ctxt dest contracted_eq_thm
58821
11e226e8a095 modernized setup;
wenzelm
parents: 56722
diff changeset
   300
  end
51956
a4d81cdebf8b better support for domains in Lifting/Transfer = replace Domainp T by the actual invariant in a transferred goal
kuncar
parents: 51955
diff changeset
   301
a4d81cdebf8b better support for domains in Lifting/Transfer = replace Domainp T by the actual invariant in a transferred goal
kuncar
parents: 51955
diff changeset
   302
a4d81cdebf8b better support for domains in Lifting/Transfer = replace Domainp T by the actual invariant in a transferred goal
kuncar
parents: 51955
diff changeset
   303
(** Replacing explicit Domainp predicates with Domainp assumptions **)
a4d81cdebf8b better support for domains in Lifting/Transfer = replace Domainp T by the actual invariant in a transferred goal
kuncar
parents: 51955
diff changeset
   304
a4d81cdebf8b better support for domains in Lifting/Transfer = replace Domainp T by the actual invariant in a transferred goal
kuncar
parents: 51955
diff changeset
   305
fun mk_Domainp_assm (T, R) =
69593
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 67664
diff changeset
   306
  HOLogic.mk_eq ((Const (\<^const_name>\<open>Domainp\<close>, Term.fastype_of T --> Term.fastype_of R) $ T), R)
51956
a4d81cdebf8b better support for domains in Lifting/Transfer = replace Domainp T by the actual invariant in a transferred goal
kuncar
parents: 51955
diff changeset
   307
a4d81cdebf8b better support for domains in Lifting/Transfer = replace Domainp T by the actual invariant in a transferred goal
kuncar
parents: 51955
diff changeset
   308
val Domainp_lemma =
a4d81cdebf8b better support for domains in Lifting/Transfer = replace Domainp T by the actual invariant in a transferred goal
kuncar
parents: 51955
diff changeset
   309
  @{lemma "(!!R. Domainp T = R ==> PROP (P R)) == PROP (P (Domainp T))"
a4d81cdebf8b better support for domains in Lifting/Transfer = replace Domainp T by the actual invariant in a transferred goal
kuncar
parents: 51955
diff changeset
   310
    by (rule, drule meta_spec,
a4d81cdebf8b better support for domains in Lifting/Transfer = replace Domainp T by the actual invariant in a transferred goal
kuncar
parents: 51955
diff changeset
   311
      erule meta_mp, rule refl, simp)}
a4d81cdebf8b better support for domains in Lifting/Transfer = replace Domainp T by the actual invariant in a transferred goal
kuncar
parents: 51955
diff changeset
   312
69593
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 67664
diff changeset
   313
fun fold_Domainp f (t as Const (\<^const_name>\<open>Domainp\<close>,_) $ (Var (_,_))) = f t
51956
a4d81cdebf8b better support for domains in Lifting/Transfer = replace Domainp T by the actual invariant in a transferred goal
kuncar
parents: 51955
diff changeset
   314
  | fold_Domainp f (t $ u) = fold_Domainp f t #> fold_Domainp f u
a4d81cdebf8b better support for domains in Lifting/Transfer = replace Domainp T by the actual invariant in a transferred goal
kuncar
parents: 51955
diff changeset
   315
  | fold_Domainp f (Abs (_, _, t)) = fold_Domainp f t
a4d81cdebf8b better support for domains in Lifting/Transfer = replace Domainp T by the actual invariant in a transferred goal
kuncar
parents: 51955
diff changeset
   316
  | fold_Domainp _ _ = I
a4d81cdebf8b better support for domains in Lifting/Transfer = replace Domainp T by the actual invariant in a transferred goal
kuncar
parents: 51955
diff changeset
   317
58821
11e226e8a095 modernized setup;
wenzelm
parents: 56722
diff changeset
   318
fun subst_terms tab t =
51956
a4d81cdebf8b better support for domains in Lifting/Transfer = replace Domainp T by the actual invariant in a transferred goal
kuncar
parents: 51955
diff changeset
   319
  let
a4d81cdebf8b better support for domains in Lifting/Transfer = replace Domainp T by the actual invariant in a transferred goal
kuncar
parents: 51955
diff changeset
   320
    val t' = Termtab.lookup tab t
a4d81cdebf8b better support for domains in Lifting/Transfer = replace Domainp T by the actual invariant in a transferred goal
kuncar
parents: 51955
diff changeset
   321
  in
a4d81cdebf8b better support for domains in Lifting/Transfer = replace Domainp T by the actual invariant in a transferred goal
kuncar
parents: 51955
diff changeset
   322
    case t' of
a4d81cdebf8b better support for domains in Lifting/Transfer = replace Domainp T by the actual invariant in a transferred goal
kuncar
parents: 51955
diff changeset
   323
      SOME t' => t'
58821
11e226e8a095 modernized setup;
wenzelm
parents: 56722
diff changeset
   324
      | NONE =>
51956
a4d81cdebf8b better support for domains in Lifting/Transfer = replace Domainp T by the actual invariant in a transferred goal
kuncar
parents: 51955
diff changeset
   325
        (case t of
a4d81cdebf8b better support for domains in Lifting/Transfer = replace Domainp T by the actual invariant in a transferred goal
kuncar
parents: 51955
diff changeset
   326
          u $ v => (subst_terms tab u) $ (subst_terms tab v)
a4d81cdebf8b better support for domains in Lifting/Transfer = replace Domainp T by the actual invariant in a transferred goal
kuncar
parents: 51955
diff changeset
   327
          | Abs (a, T, t) => Abs (a, T, subst_terms tab t)
a4d81cdebf8b better support for domains in Lifting/Transfer = replace Domainp T by the actual invariant in a transferred goal
kuncar
parents: 51955
diff changeset
   328
          | t => t)
a4d81cdebf8b better support for domains in Lifting/Transfer = replace Domainp T by the actual invariant in a transferred goal
kuncar
parents: 51955
diff changeset
   329
  end
a4d81cdebf8b better support for domains in Lifting/Transfer = replace Domainp T by the actual invariant in a transferred goal
kuncar
parents: 51955
diff changeset
   330
54742
7a86358a3c0b proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
wenzelm
parents: 53649
diff changeset
   331
fun gen_abstract_domains ctxt (dest : term -> term * (term -> term)) thm =
51956
a4d81cdebf8b better support for domains in Lifting/Transfer = replace Domainp T by the actual invariant in a transferred goal
kuncar
parents: 51955
diff changeset
   332
  let
a4d81cdebf8b better support for domains in Lifting/Transfer = replace Domainp T by the actual invariant in a transferred goal
kuncar
parents: 51955
diff changeset
   333
    val prop = Thm.prop_of thm
a4d81cdebf8b better support for domains in Lifting/Transfer = replace Domainp T by the actual invariant in a transferred goal
kuncar
parents: 51955
diff changeset
   334
    val (t, mk_prop') = dest prop
a4d81cdebf8b better support for domains in Lifting/Transfer = replace Domainp T by the actual invariant in a transferred goal
kuncar
parents: 51955
diff changeset
   335
    val Domainp_tms = rev (fold_Domainp (fn t => insert op= t) t [])
a4d81cdebf8b better support for domains in Lifting/Transfer = replace Domainp T by the actual invariant in a transferred goal
kuncar
parents: 51955
diff changeset
   336
    val Domainp_Ts = map (snd o dest_funT o snd o dest_Const o fst o dest_comb) Domainp_tms
a4d81cdebf8b better support for domains in Lifting/Transfer = replace Domainp T by the actual invariant in a transferred goal
kuncar
parents: 51955
diff changeset
   337
    val used = Term.add_free_names t []
a4d81cdebf8b better support for domains in Lifting/Transfer = replace Domainp T by the actual invariant in a transferred goal
kuncar
parents: 51955
diff changeset
   338
    val rels = map (snd o dest_comb) Domainp_tms
a4d81cdebf8b better support for domains in Lifting/Transfer = replace Domainp T by the actual invariant in a transferred goal
kuncar
parents: 51955
diff changeset
   339
    val rel_names = map (fst o fst o dest_Var) rels
a4d81cdebf8b better support for domains in Lifting/Transfer = replace Domainp T by the actual invariant in a transferred goal
kuncar
parents: 51955
diff changeset
   340
    val names = map (fn name => ("D" ^ name)) rel_names |> Name.variant_list used
a4d81cdebf8b better support for domains in Lifting/Transfer = replace Domainp T by the actual invariant in a transferred goal
kuncar
parents: 51955
diff changeset
   341
    val frees = map Free (names ~~ Domainp_Ts)
a4d81cdebf8b better support for domains in Lifting/Transfer = replace Domainp T by the actual invariant in a transferred goal
kuncar
parents: 51955
diff changeset
   342
    val prems = map (HOLogic.mk_Trueprop o mk_Domainp_assm) (rels ~~ frees);
a4d81cdebf8b better support for domains in Lifting/Transfer = replace Domainp T by the actual invariant in a transferred goal
kuncar
parents: 51955
diff changeset
   343
    val t' = subst_terms (fold Termtab.update (Domainp_tms ~~ frees) Termtab.empty) t
a4d81cdebf8b better support for domains in Lifting/Transfer = replace Domainp T by the actual invariant in a transferred goal
kuncar
parents: 51955
diff changeset
   344
    val prop1 = fold Logic.all frees (Logic.list_implies (prems, mk_prop' t'))
a4d81cdebf8b better support for domains in Lifting/Transfer = replace Domainp T by the actual invariant in a transferred goal
kuncar
parents: 51955
diff changeset
   345
    val prop2 = Logic.list_rename_params (rev names) prop1
59642
929984c529d3 clarified context;
wenzelm
parents: 59621
diff changeset
   346
    val cprop = Thm.cterm_of ctxt prop2
54742
7a86358a3c0b proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
wenzelm
parents: 53649
diff changeset
   347
    val equal_thm = Raw_Simplifier.rewrite ctxt false [Domainp_lemma] cprop
51956
a4d81cdebf8b better support for domains in Lifting/Transfer = replace Domainp T by the actual invariant in a transferred goal
kuncar
parents: 51955
diff changeset
   348
    fun forall_elim thm = Thm.forall_elim_vars (Thm.maxidx_of thm + 1) thm;
a4d81cdebf8b better support for domains in Lifting/Transfer = replace Domainp T by the actual invariant in a transferred goal
kuncar
parents: 51955
diff changeset
   349
  in
a4d81cdebf8b better support for domains in Lifting/Transfer = replace Domainp T by the actual invariant in a transferred goal
kuncar
parents: 51955
diff changeset
   350
    forall_elim (thm COMP (equal_thm COMP @{thm equal_elim_rule2}))
a4d81cdebf8b better support for domains in Lifting/Transfer = replace Domainp T by the actual invariant in a transferred goal
kuncar
parents: 51955
diff changeset
   351
  end
70316
c61b7af108a6 misc tuning and clarification, notably wrt. flow of context;
wenzelm
parents: 70159
diff changeset
   352
  handle TERM _ => thm
51956
a4d81cdebf8b better support for domains in Lifting/Transfer = replace Domainp T by the actual invariant in a transferred goal
kuncar
parents: 51955
diff changeset
   353
54742
7a86358a3c0b proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
wenzelm
parents: 53649
diff changeset
   354
fun abstract_domains_transfer ctxt thm =
51956
a4d81cdebf8b better support for domains in Lifting/Transfer = replace Domainp T by the actual invariant in a transferred goal
kuncar
parents: 51955
diff changeset
   355
  let
a4d81cdebf8b better support for domains in Lifting/Transfer = replace Domainp T by the actual invariant in a transferred goal
kuncar
parents: 51955
diff changeset
   356
    fun dest prop =
a4d81cdebf8b better support for domains in Lifting/Transfer = replace Domainp T by the actual invariant in a transferred goal
kuncar
parents: 51955
diff changeset
   357
      let
a4d81cdebf8b better support for domains in Lifting/Transfer = replace Domainp T by the actual invariant in a transferred goal
kuncar
parents: 51955
diff changeset
   358
        val prems = Logic.strip_imp_prems prop
a4d81cdebf8b better support for domains in Lifting/Transfer = replace Domainp T by the actual invariant in a transferred goal
kuncar
parents: 51955
diff changeset
   359
        val concl = HOLogic.dest_Trueprop (Logic.strip_imp_concl prop)
a4d81cdebf8b better support for domains in Lifting/Transfer = replace Domainp T by the actual invariant in a transferred goal
kuncar
parents: 51955
diff changeset
   360
        val ((rel, x), y) = apfst Term.dest_comb (Term.dest_comb concl)
a4d81cdebf8b better support for domains in Lifting/Transfer = replace Domainp T by the actual invariant in a transferred goal
kuncar
parents: 51955
diff changeset
   361
      in
a4d81cdebf8b better support for domains in Lifting/Transfer = replace Domainp T by the actual invariant in a transferred goal
kuncar
parents: 51955
diff changeset
   362
        (x, fn x' =>
a4d81cdebf8b better support for domains in Lifting/Transfer = replace Domainp T by the actual invariant in a transferred goal
kuncar
parents: 51955
diff changeset
   363
          Logic.list_implies (prems, HOLogic.mk_Trueprop (rel $ x' $ y)))
a4d81cdebf8b better support for domains in Lifting/Transfer = replace Domainp T by the actual invariant in a transferred goal
kuncar
parents: 51955
diff changeset
   364
      end
a4d81cdebf8b better support for domains in Lifting/Transfer = replace Domainp T by the actual invariant in a transferred goal
kuncar
parents: 51955
diff changeset
   365
  in
54742
7a86358a3c0b proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
wenzelm
parents: 53649
diff changeset
   366
    gen_abstract_domains ctxt dest thm
51956
a4d81cdebf8b better support for domains in Lifting/Transfer = replace Domainp T by the actual invariant in a transferred goal
kuncar
parents: 51955
diff changeset
   367
  end
a4d81cdebf8b better support for domains in Lifting/Transfer = replace Domainp T by the actual invariant in a transferred goal
kuncar
parents: 51955
diff changeset
   368
56520
3373f5d1e074 abstract Domainp in relator_domain rules => more natural statement of the rule
kuncar
parents: 56254
diff changeset
   369
fun abstract_domains_relator_domain ctxt thm =
3373f5d1e074 abstract Domainp in relator_domain rules => more natural statement of the rule
kuncar
parents: 56254
diff changeset
   370
  let
3373f5d1e074 abstract Domainp in relator_domain rules => more natural statement of the rule
kuncar
parents: 56254
diff changeset
   371
    fun dest prop =
3373f5d1e074 abstract Domainp in relator_domain rules => more natural statement of the rule
kuncar
parents: 56254
diff changeset
   372
      let
3373f5d1e074 abstract Domainp in relator_domain rules => more natural statement of the rule
kuncar
parents: 56254
diff changeset
   373
        val prems = Logic.strip_imp_prems prop
3373f5d1e074 abstract Domainp in relator_domain rules => more natural statement of the rule
kuncar
parents: 56254
diff changeset
   374
        val concl = HOLogic.dest_Trueprop (Logic.strip_imp_concl prop)
3373f5d1e074 abstract Domainp in relator_domain rules => more natural statement of the rule
kuncar
parents: 56254
diff changeset
   375
        val ((rel, x), y) = apfst Term.dest_comb (Term.dest_comb concl)
3373f5d1e074 abstract Domainp in relator_domain rules => more natural statement of the rule
kuncar
parents: 56254
diff changeset
   376
      in
3373f5d1e074 abstract Domainp in relator_domain rules => more natural statement of the rule
kuncar
parents: 56254
diff changeset
   377
        (y, fn y' =>
3373f5d1e074 abstract Domainp in relator_domain rules => more natural statement of the rule
kuncar
parents: 56254
diff changeset
   378
          Logic.list_implies (prems, HOLogic.mk_Trueprop (rel $ x $ y')))
3373f5d1e074 abstract Domainp in relator_domain rules => more natural statement of the rule
kuncar
parents: 56254
diff changeset
   379
      end
3373f5d1e074 abstract Domainp in relator_domain rules => more natural statement of the rule
kuncar
parents: 56254
diff changeset
   380
  in
3373f5d1e074 abstract Domainp in relator_domain rules => more natural statement of the rule
kuncar
parents: 56254
diff changeset
   381
    gen_abstract_domains ctxt dest thm
3373f5d1e074 abstract Domainp in relator_domain rules => more natural statement of the rule
kuncar
parents: 56254
diff changeset
   382
  end
3373f5d1e074 abstract Domainp in relator_domain rules => more natural statement of the rule
kuncar
parents: 56254
diff changeset
   383
51956
a4d81cdebf8b better support for domains in Lifting/Transfer = replace Domainp T by the actual invariant in a transferred goal
kuncar
parents: 51955
diff changeset
   384
fun detect_transfer_rules thm =
a4d81cdebf8b better support for domains in Lifting/Transfer = replace Domainp T by the actual invariant in a transferred goal
kuncar
parents: 51955
diff changeset
   385
  let
a4d81cdebf8b better support for domains in Lifting/Transfer = replace Domainp T by the actual invariant in a transferred goal
kuncar
parents: 51955
diff changeset
   386
    fun is_transfer_rule tm = case (HOLogic.dest_Trueprop tm) of
69593
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 67664
diff changeset
   387
      (Const (\<^const_name>\<open>HOL.eq\<close>, _)) $ ((Const (\<^const_name>\<open>Domainp\<close>, _)) $ _) $ _ => false
51956
a4d81cdebf8b better support for domains in Lifting/Transfer = replace Domainp T by the actual invariant in a transferred goal
kuncar
parents: 51955
diff changeset
   388
      | _ $ _ $ _ => true
a4d81cdebf8b better support for domains in Lifting/Transfer = replace Domainp T by the actual invariant in a transferred goal
kuncar
parents: 51955
diff changeset
   389
      | _ => false
a4d81cdebf8b better support for domains in Lifting/Transfer = replace Domainp T by the actual invariant in a transferred goal
kuncar
parents: 51955
diff changeset
   390
    fun safe_transfer_rule_conv ctm =
59582
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59531
diff changeset
   391
      if is_transfer_rule (Thm.term_of ctm) then safe_Rel_conv ctm else Conv.all_conv ctm
51956
a4d81cdebf8b better support for domains in Lifting/Transfer = replace Domainp T by the actual invariant in a transferred goal
kuncar
parents: 51955
diff changeset
   392
  in
a4d81cdebf8b better support for domains in Lifting/Transfer = replace Domainp T by the actual invariant in a transferred goal
kuncar
parents: 51955
diff changeset
   393
    Conv.fconv_rule (Conv.prems_conv ~1 safe_transfer_rule_conv) thm
a4d81cdebf8b better support for domains in Lifting/Transfer = replace Domainp T by the actual invariant in a transferred goal
kuncar
parents: 51955
diff changeset
   394
  end
a4d81cdebf8b better support for domains in Lifting/Transfer = replace Domainp T by the actual invariant in a transferred goal
kuncar
parents: 51955
diff changeset
   395
a4d81cdebf8b better support for domains in Lifting/Transfer = replace Domainp T by the actual invariant in a transferred goal
kuncar
parents: 51955
diff changeset
   396
(** Adding transfer domain rules **)
a4d81cdebf8b better support for domains in Lifting/Transfer = replace Domainp T by the actual invariant in a transferred goal
kuncar
parents: 51955
diff changeset
   397
70316
c61b7af108a6 misc tuning and clarification, notably wrt. flow of context;
wenzelm
parents: 70159
diff changeset
   398
fun prep_transfer_domain_thm ctxt =
c61b7af108a6 misc tuning and clarification, notably wrt. flow of context;
wenzelm
parents: 70159
diff changeset
   399
  abstract_equalities_domain ctxt o detect_transfer_rules
51956
a4d81cdebf8b better support for domains in Lifting/Transfer = replace Domainp T by the actual invariant in a transferred goal
kuncar
parents: 51955
diff changeset
   400
70316
c61b7af108a6 misc tuning and clarification, notably wrt. flow of context;
wenzelm
parents: 70159
diff changeset
   401
fun add_transfer_domain_thm thm ctxt =
c61b7af108a6 misc tuning and clarification, notably wrt. flow of context;
wenzelm
parents: 70159
diff changeset
   402
  (add_transfer_thm o prep_transfer_domain_thm (Context.proof_of ctxt)) thm ctxt
55731
66df76dd2640 rewrite composition of quotients to a more readable form in a respectfulness goal that is presented to a user
kuncar
parents: 55563
diff changeset
   403
70316
c61b7af108a6 misc tuning and clarification, notably wrt. flow of context;
wenzelm
parents: 70159
diff changeset
   404
fun del_transfer_domain_thm thm ctxt =
c61b7af108a6 misc tuning and clarification, notably wrt. flow of context;
wenzelm
parents: 70159
diff changeset
   405
  (del_transfer_thm o prep_transfer_domain_thm (Context.proof_of ctxt)) thm ctxt
49975
faf4afed009f transfer package: more flexible handling of equality relations using is_equality predicate
huffman
parents: 49625
diff changeset
   406
faf4afed009f transfer package: more flexible handling of equality relations using is_equality predicate
huffman
parents: 49625
diff changeset
   407
(** Transfer proof method **)
47325
ec6187036495 new transfer proof method
huffman
parents:
diff changeset
   408
47355
3d9d98e0f1a4 add bounded quantifier constant transfer_bforall, whose definition is unfolded after transfer
huffman
parents: 47327
diff changeset
   409
val post_simps =
47789
71a526ee569a implement transfer tactic with more scalable forward proof methods
huffman
parents: 47658
diff changeset
   410
  @{thms transfer_forall_eq [symmetric]
47355
3d9d98e0f1a4 add bounded quantifier constant transfer_bforall, whose definition is unfolded after transfer
huffman
parents: 47327
diff changeset
   411
    transfer_implies_eq [symmetric] transfer_bforall_unfold}
3d9d98e0f1a4 add bounded quantifier constant transfer_bforall, whose definition is unfolded after transfer
huffman
parents: 47327
diff changeset
   412
47356
19fb95255ec9 transfer method generalizes over free variables in goal
huffman
parents: 47355
diff changeset
   413
fun gen_frees_tac keepers ctxt = SUBGOAL (fn (t, i) =>
19fb95255ec9 transfer method generalizes over free variables in goal
huffman
parents: 47355
diff changeset
   414
  let
48065
8aa05d38299a transfer method now avoids generalizing over free variables that are known to appear in registered transfer rules
huffman
parents: 48064
diff changeset
   415
    val keepers = keepers @ get_known_frees ctxt
47356
19fb95255ec9 transfer method generalizes over free variables in goal
huffman
parents: 47355
diff changeset
   416
    val vs = rev (Term.add_frees t [])
47568
98c8b7542b72 add option to transfer method for specifying variables not to generalize over
huffman
parents: 47523
diff changeset
   417
    val vs' = filter_out (member (op =) keepers) vs
47356
19fb95255ec9 transfer method generalizes over free variables in goal
huffman
parents: 47355
diff changeset
   418
  in
19fb95255ec9 transfer method generalizes over free variables in goal
huffman
parents: 47355
diff changeset
   419
    Induct.arbitrary_tac ctxt 0 vs' i
19fb95255ec9 transfer method generalizes over free variables in goal
huffman
parents: 47355
diff changeset
   420
  end)
19fb95255ec9 transfer method generalizes over free variables in goal
huffman
parents: 47355
diff changeset
   421
47789
71a526ee569a implement transfer tactic with more scalable forward proof methods
huffman
parents: 47658
diff changeset
   422
fun mk_relT (T, U) = T --> U --> HOLogic.boolT
71a526ee569a implement transfer tactic with more scalable forward proof methods
huffman
parents: 47658
diff changeset
   423
71a526ee569a implement transfer tactic with more scalable forward proof methods
huffman
parents: 47658
diff changeset
   424
fun mk_Rel t =
71a526ee569a implement transfer tactic with more scalable forward proof methods
huffman
parents: 47658
diff changeset
   425
  let val T = fastype_of t
69593
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 67664
diff changeset
   426
  in Const (\<^const_name>\<open>Transfer.Rel\<close>, T --> T) $ t end
47789
71a526ee569a implement transfer tactic with more scalable forward proof methods
huffman
parents: 47658
diff changeset
   427
52354
acb4f932dd24 implement 'transferred' attribute for transfer package, with support for monotonicity of !!/==>
huffman
parents: 51996
diff changeset
   428
fun transfer_rule_of_terms (prj : typ * typ -> typ) ctxt tab t u =
47580
d99c883cdf2c use simpler method for preserving bound variable names in transfer tactic
huffman
parents: 47568
diff changeset
   429
  let
52354
acb4f932dd24 implement 'transferred' attribute for transfer package, with support for monotonicity of !!/==>
huffman
parents: 51996
diff changeset
   430
    (* precondition: prj(T,U) must consist of only TFrees and type "fun" *)
acb4f932dd24 implement 'transferred' attribute for transfer package, with support for monotonicity of !!/==>
huffman
parents: 51996
diff changeset
   431
    fun rel (T as Type ("fun", [T1, T2])) (U as Type ("fun", [U1, U2])) =
59642
929984c529d3 clarified context;
wenzelm
parents: 59621
diff changeset
   432
          let
929984c529d3 clarified context;
wenzelm
parents: 59621
diff changeset
   433
            val r1 = rel T1 U1
929984c529d3 clarified context;
wenzelm
parents: 59621
diff changeset
   434
            val r2 = rel T2 U2
929984c529d3 clarified context;
wenzelm
parents: 59621
diff changeset
   435
            val rT = fastype_of r1 --> fastype_of r2 --> mk_relT (T, U)
929984c529d3 clarified context;
wenzelm
parents: 59621
diff changeset
   436
          in
69593
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 67664
diff changeset
   437
            Const (\<^const_name>\<open>rel_fun\<close>, rT) $ r1 $ r2
59642
929984c529d3 clarified context;
wenzelm
parents: 59621
diff changeset
   438
          end
52354
acb4f932dd24 implement 'transferred' attribute for transfer package, with support for monotonicity of !!/==>
huffman
parents: 51996
diff changeset
   439
      | rel T U =
59642
929984c529d3 clarified context;
wenzelm
parents: 59621
diff changeset
   440
          let
929984c529d3 clarified context;
wenzelm
parents: 59621
diff changeset
   441
            val (a, _) = dest_TFree (prj (T, U))
929984c529d3 clarified context;
wenzelm
parents: 59621
diff changeset
   442
          in
929984c529d3 clarified context;
wenzelm
parents: 59621
diff changeset
   443
            Free (the (AList.lookup (op =) tab a), mk_relT (T, U))
929984c529d3 clarified context;
wenzelm
parents: 59621
diff changeset
   444
          end
47789
71a526ee569a implement transfer tactic with more scalable forward proof methods
huffman
parents: 47658
diff changeset
   445
    fun zip _ thms (Bound i) (Bound _) = (nth thms i, [])
71a526ee569a implement transfer tactic with more scalable forward proof methods
huffman
parents: 47658
diff changeset
   446
      | zip ctxt thms (Abs (x, T, t)) (Abs (y, U, u)) =
59642
929984c529d3 clarified context;
wenzelm
parents: 59621
diff changeset
   447
          let
929984c529d3 clarified context;
wenzelm
parents: 59621
diff changeset
   448
            val ([x', y'], ctxt') = Variable.variant_fixes [x, y] ctxt
929984c529d3 clarified context;
wenzelm
parents: 59621
diff changeset
   449
            val prop = mk_Rel (rel T U) $ Free (x', T) $ Free (y', U)
929984c529d3 clarified context;
wenzelm
parents: 59621
diff changeset
   450
            val cprop = Thm.cterm_of ctxt' (HOLogic.mk_Trueprop prop)
929984c529d3 clarified context;
wenzelm
parents: 59621
diff changeset
   451
            val thm0 = Thm.assume cprop
929984c529d3 clarified context;
wenzelm
parents: 59621
diff changeset
   452
            val (thm1, hyps) = zip ctxt' (thm0 :: thms) t u
929984c529d3 clarified context;
wenzelm
parents: 59621
diff changeset
   453
            val ((r1, x), y) = apfst Thm.dest_comb (Thm.dest_comb (Thm.dest_arg cprop))
929984c529d3 clarified context;
wenzelm
parents: 59621
diff changeset
   454
            val r2 = Thm.dest_fun2 (Thm.dest_arg (Thm.cprop_of thm1))
70159
57503fe1b0ff tuned signature;
wenzelm
parents: 70154
diff changeset
   455
            val (a1, (b1, _)) = apsnd Thm.dest_funT (Thm.dest_funT (Thm.ctyp_of_cterm r1))
57503fe1b0ff tuned signature;
wenzelm
parents: 70154
diff changeset
   456
            val (a2, (b2, _)) = apsnd Thm.dest_funT (Thm.dest_funT (Thm.ctyp_of_cterm r2))
59642
929984c529d3 clarified context;
wenzelm
parents: 59621
diff changeset
   457
            val tinsts = [SOME a1, SOME b1, SOME a2, SOME b2]
929984c529d3 clarified context;
wenzelm
parents: 59621
diff changeset
   458
            val insts = [SOME (Thm.dest_arg r1), SOME (Thm.dest_arg r2)]
60801
7664e0916eec tuned signature;
wenzelm
parents: 60752
diff changeset
   459
            val rule = Thm.instantiate' tinsts insts @{thm Rel_abs}
59642
929984c529d3 clarified context;
wenzelm
parents: 59621
diff changeset
   460
            val thm2 = Thm.forall_intr x (Thm.forall_intr y (Thm.implies_intr cprop thm1))
929984c529d3 clarified context;
wenzelm
parents: 59621
diff changeset
   461
          in
929984c529d3 clarified context;
wenzelm
parents: 59621
diff changeset
   462
            (thm2 COMP rule, hyps)
929984c529d3 clarified context;
wenzelm
parents: 59621
diff changeset
   463
          end
47789
71a526ee569a implement transfer tactic with more scalable forward proof methods
huffman
parents: 47658
diff changeset
   464
      | zip ctxt thms (f $ t) (g $ u) =
59642
929984c529d3 clarified context;
wenzelm
parents: 59621
diff changeset
   465
          let
929984c529d3 clarified context;
wenzelm
parents: 59621
diff changeset
   466
            val (thm1, hyps1) = zip ctxt thms f g
929984c529d3 clarified context;
wenzelm
parents: 59621
diff changeset
   467
            val (thm2, hyps2) = zip ctxt thms t u
929984c529d3 clarified context;
wenzelm
parents: 59621
diff changeset
   468
          in
929984c529d3 clarified context;
wenzelm
parents: 59621
diff changeset
   469
            (thm2 RS (thm1 RS @{thm Rel_app}), hyps1 @ hyps2)
929984c529d3 clarified context;
wenzelm
parents: 59621
diff changeset
   470
          end
52354
acb4f932dd24 implement 'transferred' attribute for transfer package, with support for monotonicity of !!/==>
huffman
parents: 51996
diff changeset
   471
      | zip _ _ t u =
59642
929984c529d3 clarified context;
wenzelm
parents: 59621
diff changeset
   472
          let
929984c529d3 clarified context;
wenzelm
parents: 59621
diff changeset
   473
            val T = fastype_of t
929984c529d3 clarified context;
wenzelm
parents: 59621
diff changeset
   474
            val U = fastype_of u
929984c529d3 clarified context;
wenzelm
parents: 59621
diff changeset
   475
            val prop = mk_Rel (rel T U) $ t $ u
929984c529d3 clarified context;
wenzelm
parents: 59621
diff changeset
   476
            val cprop = Thm.cterm_of ctxt (HOLogic.mk_Trueprop prop)
929984c529d3 clarified context;
wenzelm
parents: 59621
diff changeset
   477
          in
929984c529d3 clarified context;
wenzelm
parents: 59621
diff changeset
   478
            (Thm.assume cprop, [cprop])
929984c529d3 clarified context;
wenzelm
parents: 59621
diff changeset
   479
          end
47789
71a526ee569a implement transfer tactic with more scalable forward proof methods
huffman
parents: 47658
diff changeset
   480
    val r = mk_Rel (rel (fastype_of t) (fastype_of u))
71a526ee569a implement transfer tactic with more scalable forward proof methods
huffman
parents: 47658
diff changeset
   481
    val goal = HOLogic.mk_Trueprop (r $ t $ u)
59642
929984c529d3 clarified context;
wenzelm
parents: 59621
diff changeset
   482
    val rename = Thm.trivial (Thm.cterm_of ctxt goal)
47789
71a526ee569a implement transfer tactic with more scalable forward proof methods
huffman
parents: 47658
diff changeset
   483
    val (thm, hyps) = zip ctxt [] t u
47580
d99c883cdf2c use simpler method for preserving bound variable names in transfer tactic
huffman
parents: 47568
diff changeset
   484
  in
47789
71a526ee569a implement transfer tactic with more scalable forward proof methods
huffman
parents: 47658
diff changeset
   485
    Drule.implies_intr_list hyps (thm RS rename)
71a526ee569a implement transfer tactic with more scalable forward proof methods
huffman
parents: 47658
diff changeset
   486
  end
47580
d99c883cdf2c use simpler method for preserving bound variable names in transfer tactic
huffman
parents: 47568
diff changeset
   487
52354
acb4f932dd24 implement 'transferred' attribute for transfer package, with support for monotonicity of !!/==>
huffman
parents: 51996
diff changeset
   488
(* create a lambda term of the same shape as the given term *)
70316
c61b7af108a6 misc tuning and clarification, notably wrt. flow of context;
wenzelm
parents: 70159
diff changeset
   489
fun skeleton is_atom =
47325
ec6187036495 new transfer proof method
huffman
parents:
diff changeset
   490
  let
48066
c6783c9b87bf transfer method now handles transfer rules for compound terms, e.g. locale-defined constants with hidden parameters
huffman
parents: 48065
diff changeset
   491
    fun dummy ctxt =
70316
c61b7af108a6 misc tuning and clarification, notably wrt. flow of context;
wenzelm
parents: 70159
diff changeset
   492
      let val (c, ctxt') = yield_singleton Variable.variant_fixes "a" ctxt
c61b7af108a6 misc tuning and clarification, notably wrt. flow of context;
wenzelm
parents: 70159
diff changeset
   493
      in (Free (c, dummyT), ctxt') end
c61b7af108a6 misc tuning and clarification, notably wrt. flow of context;
wenzelm
parents: 70159
diff changeset
   494
    fun skel (Bound i) ctxt = (Bound i, ctxt)
c61b7af108a6 misc tuning and clarification, notably wrt. flow of context;
wenzelm
parents: 70159
diff changeset
   495
      | skel (Abs (x, _, t)) ctxt =
c61b7af108a6 misc tuning and clarification, notably wrt. flow of context;
wenzelm
parents: 70159
diff changeset
   496
          let val (t', ctxt) = skel t ctxt
c61b7af108a6 misc tuning and clarification, notably wrt. flow of context;
wenzelm
parents: 70159
diff changeset
   497
          in (Abs (x, dummyT, t'), ctxt) end
c61b7af108a6 misc tuning and clarification, notably wrt. flow of context;
wenzelm
parents: 70159
diff changeset
   498
      | skel (tu as t $ u) ctxt =
c61b7af108a6 misc tuning and clarification, notably wrt. flow of context;
wenzelm
parents: 70159
diff changeset
   499
          if is_atom tu andalso not (Term.is_open tu) then dummy ctxt
c61b7af108a6 misc tuning and clarification, notably wrt. flow of context;
wenzelm
parents: 70159
diff changeset
   500
          else
c61b7af108a6 misc tuning and clarification, notably wrt. flow of context;
wenzelm
parents: 70159
diff changeset
   501
            let
c61b7af108a6 misc tuning and clarification, notably wrt. flow of context;
wenzelm
parents: 70159
diff changeset
   502
              val (t', ctxt) = skel t ctxt
c61b7af108a6 misc tuning and clarification, notably wrt. flow of context;
wenzelm
parents: 70159
diff changeset
   503
              val (u', ctxt) = skel u ctxt
c61b7af108a6 misc tuning and clarification, notably wrt. flow of context;
wenzelm
parents: 70159
diff changeset
   504
            in (t' $ u', ctxt) end
c61b7af108a6 misc tuning and clarification, notably wrt. flow of context;
wenzelm
parents: 70159
diff changeset
   505
      | skel _ ctxt = dummy ctxt
52354
acb4f932dd24 implement 'transferred' attribute for transfer package, with support for monotonicity of !!/==>
huffman
parents: 51996
diff changeset
   506
  in
70316
c61b7af108a6 misc tuning and clarification, notably wrt. flow of context;
wenzelm
parents: 70159
diff changeset
   507
    fn ctxt => fn t =>
c61b7af108a6 misc tuning and clarification, notably wrt. flow of context;
wenzelm
parents: 70159
diff changeset
   508
      fst (skel t ctxt) |> Syntax.check_term ctxt  (* FIXME avoid syntax operation!? *)
52354
acb4f932dd24 implement 'transferred' attribute for transfer package, with support for monotonicity of !!/==>
huffman
parents: 51996
diff changeset
   509
  end
acb4f932dd24 implement 'transferred' attribute for transfer package, with support for monotonicity of !!/==>
huffman
parents: 51996
diff changeset
   510
acb4f932dd24 implement 'transferred' attribute for transfer package, with support for monotonicity of !!/==>
huffman
parents: 51996
diff changeset
   511
(** Monotonicity analysis **)
acb4f932dd24 implement 'transferred' attribute for transfer package, with support for monotonicity of !!/==>
huffman
parents: 51996
diff changeset
   512
acb4f932dd24 implement 'transferred' attribute for transfer package, with support for monotonicity of !!/==>
huffman
parents: 51996
diff changeset
   513
(* TODO: Put extensible table in theory data *)
acb4f932dd24 implement 'transferred' attribute for transfer package, with support for monotonicity of !!/==>
huffman
parents: 51996
diff changeset
   514
val monotab =
acb4f932dd24 implement 'transferred' attribute for transfer package, with support for monotonicity of !!/==>
huffman
parents: 51996
diff changeset
   515
  Symtab.make
69593
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 67664
diff changeset
   516
    [(\<^const_name>\<open>transfer_implies\<close>, [~1, 1]),
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 67664
diff changeset
   517
     (\<^const_name>\<open>transfer_forall\<close>, [1])(*,
52354
acb4f932dd24 implement 'transferred' attribute for transfer package, with support for monotonicity of !!/==>
huffman
parents: 51996
diff changeset
   518
     (@{const_name implies}, [~1, 1]),
acb4f932dd24 implement 'transferred' attribute for transfer package, with support for monotonicity of !!/==>
huffman
parents: 51996
diff changeset
   519
     (@{const_name All}, [1])*)]
acb4f932dd24 implement 'transferred' attribute for transfer package, with support for monotonicity of !!/==>
huffman
parents: 51996
diff changeset
   520
acb4f932dd24 implement 'transferred' attribute for transfer package, with support for monotonicity of !!/==>
huffman
parents: 51996
diff changeset
   521
(*
acb4f932dd24 implement 'transferred' attribute for transfer package, with support for monotonicity of !!/==>
huffman
parents: 51996
diff changeset
   522
Function bool_insts determines the set of boolean-relation variables
acb4f932dd24 implement 'transferred' attribute for transfer package, with support for monotonicity of !!/==>
huffman
parents: 51996
diff changeset
   523
that can be instantiated to implies, rev_implies, or iff.
acb4f932dd24 implement 'transferred' attribute for transfer package, with support for monotonicity of !!/==>
huffman
parents: 51996
diff changeset
   524
acb4f932dd24 implement 'transferred' attribute for transfer package, with support for monotonicity of !!/==>
huffman
parents: 51996
diff changeset
   525
Invariants: bool_insts p (t, u) requires that
acb4f932dd24 implement 'transferred' attribute for transfer package, with support for monotonicity of !!/==>
huffman
parents: 51996
diff changeset
   526
  u :: _ => _ => ... => bool, and
acb4f932dd24 implement 'transferred' attribute for transfer package, with support for monotonicity of !!/==>
huffman
parents: 51996
diff changeset
   527
  t is a skeleton of u
acb4f932dd24 implement 'transferred' attribute for transfer package, with support for monotonicity of !!/==>
huffman
parents: 51996
diff changeset
   528
*)
acb4f932dd24 implement 'transferred' attribute for transfer package, with support for monotonicity of !!/==>
huffman
parents: 51996
diff changeset
   529
fun bool_insts p (t, u) =
acb4f932dd24 implement 'transferred' attribute for transfer package, with support for monotonicity of !!/==>
huffman
parents: 51996
diff changeset
   530
  let
acb4f932dd24 implement 'transferred' attribute for transfer package, with support for monotonicity of !!/==>
huffman
parents: 51996
diff changeset
   531
    fun strip2 (t1 $ t2, u1 $ u2, tus) =
acb4f932dd24 implement 'transferred' attribute for transfer package, with support for monotonicity of !!/==>
huffman
parents: 51996
diff changeset
   532
        strip2 (t1, u1, (t2, u2) :: tus)
acb4f932dd24 implement 'transferred' attribute for transfer package, with support for monotonicity of !!/==>
huffman
parents: 51996
diff changeset
   533
      | strip2 x = x
acb4f932dd24 implement 'transferred' attribute for transfer package, with support for monotonicity of !!/==>
huffman
parents: 51996
diff changeset
   534
    fun or3 ((a, b, c), (x, y, z)) = (a orelse x, b orelse y, c orelse z)
acb4f932dd24 implement 'transferred' attribute for transfer package, with support for monotonicity of !!/==>
huffman
parents: 51996
diff changeset
   535
    fun go Ts p (Abs (_, T, t), Abs (_, _, u)) tab = go (T :: Ts) p (t, u) tab
acb4f932dd24 implement 'transferred' attribute for transfer package, with support for monotonicity of !!/==>
huffman
parents: 51996
diff changeset
   536
      | go Ts p (t, u) tab =
acb4f932dd24 implement 'transferred' attribute for transfer package, with support for monotonicity of !!/==>
huffman
parents: 51996
diff changeset
   537
        let
acb4f932dd24 implement 'transferred' attribute for transfer package, with support for monotonicity of !!/==>
huffman
parents: 51996
diff changeset
   538
          val (a, _) = dest_TFree (Term.body_type (Term.fastype_of1 (Ts, t)))
acb4f932dd24 implement 'transferred' attribute for transfer package, with support for monotonicity of !!/==>
huffman
parents: 51996
diff changeset
   539
          val (_, tf, tus) = strip2 (t, u, [])
acb4f932dd24 implement 'transferred' attribute for transfer package, with support for monotonicity of !!/==>
huffman
parents: 51996
diff changeset
   540
          val ps_opt = case tf of Const (c, _) => Symtab.lookup monotab c | _ => NONE
acb4f932dd24 implement 'transferred' attribute for transfer package, with support for monotonicity of !!/==>
huffman
parents: 51996
diff changeset
   541
          val tab1 =
acb4f932dd24 implement 'transferred' attribute for transfer package, with support for monotonicity of !!/==>
huffman
parents: 51996
diff changeset
   542
            case ps_opt of
acb4f932dd24 implement 'transferred' attribute for transfer package, with support for monotonicity of !!/==>
huffman
parents: 51996
diff changeset
   543
              SOME ps =>
acb4f932dd24 implement 'transferred' attribute for transfer package, with support for monotonicity of !!/==>
huffman
parents: 51996
diff changeset
   544
              let
acb4f932dd24 implement 'transferred' attribute for transfer package, with support for monotonicity of !!/==>
huffman
parents: 51996
diff changeset
   545
                val ps' = map (fn x => p * x) (take (length tus) ps)
acb4f932dd24 implement 'transferred' attribute for transfer package, with support for monotonicity of !!/==>
huffman
parents: 51996
diff changeset
   546
              in
acb4f932dd24 implement 'transferred' attribute for transfer package, with support for monotonicity of !!/==>
huffman
parents: 51996
diff changeset
   547
                fold I (map2 (go Ts) ps' tus) tab
acb4f932dd24 implement 'transferred' attribute for transfer package, with support for monotonicity of !!/==>
huffman
parents: 51996
diff changeset
   548
              end
acb4f932dd24 implement 'transferred' attribute for transfer package, with support for monotonicity of !!/==>
huffman
parents: 51996
diff changeset
   549
            | NONE => tab
acb4f932dd24 implement 'transferred' attribute for transfer package, with support for monotonicity of !!/==>
huffman
parents: 51996
diff changeset
   550
          val tab2 = Symtab.make [(a, (p >= 0, p <= 0, is_none ps_opt))]
acb4f932dd24 implement 'transferred' attribute for transfer package, with support for monotonicity of !!/==>
huffman
parents: 51996
diff changeset
   551
        in
acb4f932dd24 implement 'transferred' attribute for transfer package, with support for monotonicity of !!/==>
huffman
parents: 51996
diff changeset
   552
          Symtab.join (K or3) (tab1, tab2)
acb4f932dd24 implement 'transferred' attribute for transfer package, with support for monotonicity of !!/==>
huffman
parents: 51996
diff changeset
   553
        end
acb4f932dd24 implement 'transferred' attribute for transfer package, with support for monotonicity of !!/==>
huffman
parents: 51996
diff changeset
   554
    val tab = go [] p (t, u) Symtab.empty
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69593
diff changeset
   555
    fun f (a, (true, false, false)) = SOME (a, \<^const>\<open>implies\<close>)
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69593
diff changeset
   556
      | f (a, (false, true, false)) = SOME (a, \<^const>\<open>rev_implies\<close>)
52354
acb4f932dd24 implement 'transferred' attribute for transfer package, with support for monotonicity of !!/==>
huffman
parents: 51996
diff changeset
   557
      | f (a, (true, true, _))      = SOME (a, HOLogic.eq_const HOLogic.boolT)
acb4f932dd24 implement 'transferred' attribute for transfer package, with support for monotonicity of !!/==>
huffman
parents: 51996
diff changeset
   558
      | f _                         = NONE
acb4f932dd24 implement 'transferred' attribute for transfer package, with support for monotonicity of !!/==>
huffman
parents: 51996
diff changeset
   559
  in
acb4f932dd24 implement 'transferred' attribute for transfer package, with support for monotonicity of !!/==>
huffman
parents: 51996
diff changeset
   560
    map_filter f (Symtab.dest tab)
acb4f932dd24 implement 'transferred' attribute for transfer package, with support for monotonicity of !!/==>
huffman
parents: 51996
diff changeset
   561
  end
acb4f932dd24 implement 'transferred' attribute for transfer package, with support for monotonicity of !!/==>
huffman
parents: 51996
diff changeset
   562
53145
2fb458aceb78 delete corresponding compound lhs and rhs when a transfer rule is deleted; tuned
kuncar
parents: 53144
diff changeset
   563
fun retrieve_terms t net = map fst (Item_Net.retrieve net t)
58821
11e226e8a095 modernized setup;
wenzelm
parents: 56722
diff changeset
   564
11e226e8a095 modernized setup;
wenzelm
parents: 56722
diff changeset
   565
fun matches_list ctxt term =
53145
2fb458aceb78 delete corresponding compound lhs and rhs when a transfer rule is deleted; tuned
kuncar
parents: 53144
diff changeset
   566
  is_some o find_first (fn pat => Pattern.matches (Proof_Context.theory_of ctxt) (pat, term))
53131
701360318565 double check that lhs or rhs really matches a subterm in a goal when creating a hole in a skeleton (Net.net does only rough matching)
kuncar
parents: 53042
diff changeset
   567
70316
c61b7af108a6 misc tuning and clarification, notably wrt. flow of context;
wenzelm
parents: 70159
diff changeset
   568
fun transfer_rule_of_term ctxt equiv t =
52354
acb4f932dd24 implement 'transferred' attribute for transfer package, with support for monotonicity of !!/==>
huffman
parents: 51996
diff changeset
   569
  let
acb4f932dd24 implement 'transferred' attribute for transfer package, with support for monotonicity of !!/==>
huffman
parents: 51996
diff changeset
   570
    val compound_rhs = get_compound_rhs ctxt
53145
2fb458aceb78 delete corresponding compound lhs and rhs when a transfer rule is deleted; tuned
kuncar
parents: 53144
diff changeset
   571
    fun is_rhs t = compound_rhs |> retrieve_terms t |> matches_list ctxt t
52354
acb4f932dd24 implement 'transferred' attribute for transfer package, with support for monotonicity of !!/==>
huffman
parents: 51996
diff changeset
   572
    val s = skeleton is_rhs ctxt t
47789
71a526ee569a implement transfer tactic with more scalable forward proof methods
huffman
parents: 47658
diff changeset
   573
    val frees = map fst (Term.add_frees s [])
71a526ee569a implement transfer tactic with more scalable forward proof methods
huffman
parents: 47658
diff changeset
   574
    val tfrees = map fst (Term.add_tfrees s [])
71a526ee569a implement transfer tactic with more scalable forward proof methods
huffman
parents: 47658
diff changeset
   575
    fun prep a = "R" ^ Library.unprefix "'" a
71a526ee569a implement transfer tactic with more scalable forward proof methods
huffman
parents: 47658
diff changeset
   576
    val (rnames, ctxt') = Variable.variant_fixes (map prep tfrees) ctxt
52354
acb4f932dd24 implement 'transferred' attribute for transfer package, with support for monotonicity of !!/==>
huffman
parents: 51996
diff changeset
   577
    val tab = tfrees ~~ rnames
acb4f932dd24 implement 'transferred' attribute for transfer package, with support for monotonicity of !!/==>
huffman
parents: 51996
diff changeset
   578
    fun prep a = the (AList.lookup (op =) tab a)
acb4f932dd24 implement 'transferred' attribute for transfer package, with support for monotonicity of !!/==>
huffman
parents: 51996
diff changeset
   579
    val thm = transfer_rule_of_terms fst ctxt' tab s t
acb4f932dd24 implement 'transferred' attribute for transfer package, with support for monotonicity of !!/==>
huffman
parents: 51996
diff changeset
   580
    val binsts = bool_insts (if equiv then 0 else 1) (s, t)
acb4f932dd24 implement 'transferred' attribute for transfer package, with support for monotonicity of !!/==>
huffman
parents: 51996
diff changeset
   581
    val idx = Thm.maxidx_of thm + 1
69593
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 67664
diff changeset
   582
    fun tinst (a, _) = (((a, idx), \<^sort>\<open>type\<close>), \<^ctyp>\<open>bool\<close>)
60642
48dd1cefb4ae simplified Thm.instantiate and derivatives: the LHS refers to non-certified variables -- this merely serves as index into already certified structures (or is ignored);
wenzelm
parents: 60367
diff changeset
   583
    fun inst (a, t) =
69593
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 67664
diff changeset
   584
      ((Name.clean_index (prep a, idx), \<^typ>\<open>bool \<Rightarrow> bool \<Rightarrow> bool\<close>), Thm.cterm_of ctxt' t)
47789
71a526ee569a implement transfer tactic with more scalable forward proof methods
huffman
parents: 47658
diff changeset
   585
  in
52354
acb4f932dd24 implement 'transferred' attribute for transfer package, with support for monotonicity of !!/==>
huffman
parents: 51996
diff changeset
   586
    thm
70316
c61b7af108a6 misc tuning and clarification, notably wrt. flow of context;
wenzelm
parents: 70159
diff changeset
   587
    |> Thm.generalize (tfrees, rnames @ frees) idx
c61b7af108a6 misc tuning and clarification, notably wrt. flow of context;
wenzelm
parents: 70159
diff changeset
   588
    |> Thm.instantiate (map tinst binsts, map inst binsts)
52354
acb4f932dd24 implement 'transferred' attribute for transfer package, with support for monotonicity of !!/==>
huffman
parents: 51996
diff changeset
   589
  end
acb4f932dd24 implement 'transferred' attribute for transfer package, with support for monotonicity of !!/==>
huffman
parents: 51996
diff changeset
   590
70316
c61b7af108a6 misc tuning and clarification, notably wrt. flow of context;
wenzelm
parents: 70159
diff changeset
   591
fun transfer_rule_of_lhs ctxt t =
52354
acb4f932dd24 implement 'transferred' attribute for transfer package, with support for monotonicity of !!/==>
huffman
parents: 51996
diff changeset
   592
  let
acb4f932dd24 implement 'transferred' attribute for transfer package, with support for monotonicity of !!/==>
huffman
parents: 51996
diff changeset
   593
    val compound_lhs = get_compound_lhs ctxt
53145
2fb458aceb78 delete corresponding compound lhs and rhs when a transfer rule is deleted; tuned
kuncar
parents: 53144
diff changeset
   594
    fun is_lhs t = compound_lhs |> retrieve_terms t |> matches_list ctxt t
52354
acb4f932dd24 implement 'transferred' attribute for transfer package, with support for monotonicity of !!/==>
huffman
parents: 51996
diff changeset
   595
    val s = skeleton is_lhs ctxt t
acb4f932dd24 implement 'transferred' attribute for transfer package, with support for monotonicity of !!/==>
huffman
parents: 51996
diff changeset
   596
    val frees = map fst (Term.add_frees s [])
acb4f932dd24 implement 'transferred' attribute for transfer package, with support for monotonicity of !!/==>
huffman
parents: 51996
diff changeset
   597
    val tfrees = map fst (Term.add_tfrees s [])
acb4f932dd24 implement 'transferred' attribute for transfer package, with support for monotonicity of !!/==>
huffman
parents: 51996
diff changeset
   598
    fun prep a = "R" ^ Library.unprefix "'" a
acb4f932dd24 implement 'transferred' attribute for transfer package, with support for monotonicity of !!/==>
huffman
parents: 51996
diff changeset
   599
    val (rnames, ctxt') = Variable.variant_fixes (map prep tfrees) ctxt
acb4f932dd24 implement 'transferred' attribute for transfer package, with support for monotonicity of !!/==>
huffman
parents: 51996
diff changeset
   600
    val tab = tfrees ~~ rnames
acb4f932dd24 implement 'transferred' attribute for transfer package, with support for monotonicity of !!/==>
huffman
parents: 51996
diff changeset
   601
    fun prep a = the (AList.lookup (op =) tab a)
acb4f932dd24 implement 'transferred' attribute for transfer package, with support for monotonicity of !!/==>
huffman
parents: 51996
diff changeset
   602
    val thm = transfer_rule_of_terms snd ctxt' tab t s
acb4f932dd24 implement 'transferred' attribute for transfer package, with support for monotonicity of !!/==>
huffman
parents: 51996
diff changeset
   603
    val binsts = bool_insts 1 (s, t)
acb4f932dd24 implement 'transferred' attribute for transfer package, with support for monotonicity of !!/==>
huffman
parents: 51996
diff changeset
   604
    val idx = Thm.maxidx_of thm + 1
69593
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 67664
diff changeset
   605
    fun tinst (a, _) = (((a, idx), \<^sort>\<open>type\<close>), \<^ctyp>\<open>bool\<close>)
60642
48dd1cefb4ae simplified Thm.instantiate and derivatives: the LHS refers to non-certified variables -- this merely serves as index into already certified structures (or is ignored);
wenzelm
parents: 60367
diff changeset
   606
    fun inst (a, t) =
69593
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 67664
diff changeset
   607
      ((Name.clean_index (prep a, idx), \<^typ>\<open>bool \<Rightarrow> bool \<Rightarrow> bool\<close>), Thm.cterm_of ctxt' t)
52354
acb4f932dd24 implement 'transferred' attribute for transfer package, with support for monotonicity of !!/==>
huffman
parents: 51996
diff changeset
   608
  in
acb4f932dd24 implement 'transferred' attribute for transfer package, with support for monotonicity of !!/==>
huffman
parents: 51996
diff changeset
   609
    thm
70316
c61b7af108a6 misc tuning and clarification, notably wrt. flow of context;
wenzelm
parents: 70159
diff changeset
   610
    |> Thm.generalize (tfrees, rnames @ frees) idx
c61b7af108a6 misc tuning and clarification, notably wrt. flow of context;
wenzelm
parents: 70159
diff changeset
   611
    |> Thm.instantiate (map tinst binsts, map inst binsts)
47789
71a526ee569a implement transfer tactic with more scalable forward proof methods
huffman
parents: 47658
diff changeset
   612
  end
71a526ee569a implement transfer tactic with more scalable forward proof methods
huffman
parents: 47658
diff changeset
   613
59531
7c433cca8fe0 proper context;
wenzelm
parents: 59498
diff changeset
   614
fun eq_rules_tac ctxt eq_rules =
7c433cca8fe0 proper context;
wenzelm
parents: 59498
diff changeset
   615
  TRY o REPEAT_ALL_NEW (resolve_tac ctxt eq_rules)
60752
b48830b670a1 prefer tactics with explicit context;
wenzelm
parents: 60642
diff changeset
   616
  THEN_ALL_NEW resolve_tac ctxt @{thms is_equality_eq}
55731
66df76dd2640 rewrite composition of quotients to a more readable form in a respectfulness goal that is presented to a user
kuncar
parents: 55563
diff changeset
   617
59531
7c433cca8fe0 proper context;
wenzelm
parents: 59498
diff changeset
   618
fun eq_tac ctxt = eq_rules_tac ctxt (get_relator_eq_raw ctxt)
55731
66df76dd2640 rewrite composition of quotients to a more readable form in a respectfulness goal that is presented to a user
kuncar
parents: 55563
diff changeset
   619
59531
7c433cca8fe0 proper context;
wenzelm
parents: 59498
diff changeset
   620
fun transfer_step_tac ctxt =
61366
3933c0d7edc9 right parenthesization
kuncar
parents: 61075
diff changeset
   621
  REPEAT_ALL_NEW (resolve_tac ctxt (get_transfer_raw ctxt))
3933c0d7edc9 right parenthesization
kuncar
parents: 61075
diff changeset
   622
  THEN_ALL_NEW (DETERM o eq_rules_tac ctxt (get_relator_eq_raw ctxt))
51437
8739f8abbecb fixing transfer tactic - unfold fully identity relation by using relator_eq
kuncar
parents: 51374
diff changeset
   623
61367
46af4f577c7e new methods for debugging transfer and transfer_prover
kuncar
parents: 61366
diff changeset
   624
fun transfer_start_tac equiv ctxt i =
46af4f577c7e new methods for debugging transfer and transfer_prover
kuncar
parents: 61366
diff changeset
   625
    let
46af4f577c7e new methods for debugging transfer and transfer_prover
kuncar
parents: 61366
diff changeset
   626
      val pre_simps = @{thms transfer_forall_eq transfer_implies_eq}
46af4f577c7e new methods for debugging transfer and transfer_prover
kuncar
parents: 61366
diff changeset
   627
      val start_rule =
46af4f577c7e new methods for debugging transfer and transfer_prover
kuncar
parents: 61366
diff changeset
   628
        if equiv then @{thm transfer_start} else @{thm transfer_start'}
46af4f577c7e new methods for debugging transfer and transfer_prover
kuncar
parents: 61366
diff changeset
   629
      val err_msg = "Transfer failed to convert goal to an object-logic formula"
46af4f577c7e new methods for debugging transfer and transfer_prover
kuncar
parents: 61366
diff changeset
   630
      fun main_tac (t, i) =
46af4f577c7e new methods for debugging transfer and transfer_prover
kuncar
parents: 61366
diff changeset
   631
        resolve_tac ctxt [start_rule] i THEN
46af4f577c7e new methods for debugging transfer and transfer_prover
kuncar
parents: 61366
diff changeset
   632
        (resolve_tac ctxt [transfer_rule_of_term ctxt equiv (HOLogic.dest_Trueprop t)]) (i + 1)
46af4f577c7e new methods for debugging transfer and transfer_prover
kuncar
parents: 61366
diff changeset
   633
          handle TERM (_, ts) => raise TERM (err_msg, ts)
46af4f577c7e new methods for debugging transfer and transfer_prover
kuncar
parents: 61366
diff changeset
   634
    in
46af4f577c7e new methods for debugging transfer and transfer_prover
kuncar
parents: 61366
diff changeset
   635
      EVERY
46af4f577c7e new methods for debugging transfer and transfer_prover
kuncar
parents: 61366
diff changeset
   636
        [rewrite_goal_tac ctxt pre_simps i THEN
46af4f577c7e new methods for debugging transfer and transfer_prover
kuncar
parents: 61366
diff changeset
   637
         SUBGOAL main_tac i]
46af4f577c7e new methods for debugging transfer and transfer_prover
kuncar
parents: 61366
diff changeset
   638
    end;
47325
ec6187036495 new transfer proof method
huffman
parents:
diff changeset
   639
61367
46af4f577c7e new methods for debugging transfer and transfer_prover
kuncar
parents: 61366
diff changeset
   640
fun transfer_prover_start_tac ctxt = SUBGOAL (fn (t, i) =>
47325
ec6187036495 new transfer proof method
huffman
parents:
diff changeset
   641
  let
47789
71a526ee569a implement transfer tactic with more scalable forward proof methods
huffman
parents: 47658
diff changeset
   642
    val rhs = (snd o Term.dest_comb o HOLogic.dest_Trueprop) t
52354
acb4f932dd24 implement 'transferred' attribute for transfer package, with support for monotonicity of !!/==>
huffman
parents: 51996
diff changeset
   643
    val rule1 = transfer_rule_of_term ctxt false rhs
64434
af5235830c16 always expand equalities in the transfer relation in transfer_prover (cf. 0a7c97c76f46)
kuncar
parents: 62958
diff changeset
   644
    val expand_eqs_in_rel_conv = transfer_rel_conv (unfold_relator_eqs_conv ctxt)
47325
ec6187036495 new transfer proof method
huffman
parents:
diff changeset
   645
  in
ec6187036495 new transfer proof method
huffman
parents:
diff changeset
   646
    EVERY
47789
71a526ee569a implement transfer tactic with more scalable forward proof methods
huffman
parents: 47658
diff changeset
   647
      [CONVERSION prep_conv i,
64434
af5235830c16 always expand equalities in the transfer relation in transfer_prover (cf. 0a7c97c76f46)
kuncar
parents: 62958
diff changeset
   648
       CONVERSION expand_eqs_in_rel_conv i,
60752
b48830b670a1 prefer tactics with explicit context;
wenzelm
parents: 60642
diff changeset
   649
       resolve_tac ctxt @{thms transfer_prover_start} i,
64434
af5235830c16 always expand equalities in the transfer relation in transfer_prover (cf. 0a7c97c76f46)
kuncar
parents: 62958
diff changeset
   650
       resolve_tac ctxt [rule1] (i + 1)]
61367
46af4f577c7e new methods for debugging transfer and transfer_prover
kuncar
parents: 61366
diff changeset
   651
  end);
46af4f577c7e new methods for debugging transfer and transfer_prover
kuncar
parents: 61366
diff changeset
   652
46af4f577c7e new methods for debugging transfer and transfer_prover
kuncar
parents: 61366
diff changeset
   653
fun transfer_end_tac ctxt i =
46af4f577c7e new methods for debugging transfer and transfer_prover
kuncar
parents: 61366
diff changeset
   654
  let
46af4f577c7e new methods for debugging transfer and transfer_prover
kuncar
parents: 61366
diff changeset
   655
    val post_simps = @{thms transfer_forall_eq [symmetric]
46af4f577c7e new methods for debugging transfer and transfer_prover
kuncar
parents: 61366
diff changeset
   656
      transfer_implies_eq [symmetric] transfer_bforall_unfold}
46af4f577c7e new methods for debugging transfer and transfer_prover
kuncar
parents: 61366
diff changeset
   657
  in
46af4f577c7e new methods for debugging transfer and transfer_prover
kuncar
parents: 61366
diff changeset
   658
    EVERY [rewrite_goal_tac ctxt post_simps i,
46af4f577c7e new methods for debugging transfer and transfer_prover
kuncar
parents: 61366
diff changeset
   659
           Goal.norm_hhf_tac ctxt i]
46af4f577c7e new methods for debugging transfer and transfer_prover
kuncar
parents: 61366
diff changeset
   660
  end;
46af4f577c7e new methods for debugging transfer and transfer_prover
kuncar
parents: 61366
diff changeset
   661
46af4f577c7e new methods for debugging transfer and transfer_prover
kuncar
parents: 61366
diff changeset
   662
fun transfer_prover_end_tac ctxt i = resolve_tac ctxt @{thms refl} i
46af4f577c7e new methods for debugging transfer and transfer_prover
kuncar
parents: 61366
diff changeset
   663
46af4f577c7e new methods for debugging transfer and transfer_prover
kuncar
parents: 61366
diff changeset
   664
local 
46af4f577c7e new methods for debugging transfer and transfer_prover
kuncar
parents: 61366
diff changeset
   665
  infix 1 THEN_ALL_BUT_FIRST_NEW
46af4f577c7e new methods for debugging transfer and transfer_prover
kuncar
parents: 61366
diff changeset
   666
  fun (tac1 THEN_ALL_BUT_FIRST_NEW tac2) i st =
46af4f577c7e new methods for debugging transfer and transfer_prover
kuncar
parents: 61366
diff changeset
   667
    st |> (tac1 i THEN (fn st' =>
46af4f577c7e new methods for debugging transfer and transfer_prover
kuncar
parents: 61366
diff changeset
   668
      Seq.INTERVAL tac2 (i + 1) (i + Thm.nprems_of st' - Thm.nprems_of st) st'));
46af4f577c7e new methods for debugging transfer and transfer_prover
kuncar
parents: 61366
diff changeset
   669
in
46af4f577c7e new methods for debugging transfer and transfer_prover
kuncar
parents: 61366
diff changeset
   670
  fun transfer_tac equiv ctxt i =
46af4f577c7e new methods for debugging transfer and transfer_prover
kuncar
parents: 61366
diff changeset
   671
    let
46af4f577c7e new methods for debugging transfer and transfer_prover
kuncar
parents: 61366
diff changeset
   672
      val rules = get_transfer_raw ctxt
46af4f577c7e new methods for debugging transfer and transfer_prover
kuncar
parents: 61366
diff changeset
   673
      val eq_rules = get_relator_eq_raw ctxt
46af4f577c7e new methods for debugging transfer and transfer_prover
kuncar
parents: 61366
diff changeset
   674
      (* allow unsolved subgoals only for standard transfer method, not for transfer' *)
46af4f577c7e new methods for debugging transfer and transfer_prover
kuncar
parents: 61366
diff changeset
   675
      val end_tac = if equiv then K all_tac else K no_tac
46af4f577c7e new methods for debugging transfer and transfer_prover
kuncar
parents: 61366
diff changeset
   676
  
46af4f577c7e new methods for debugging transfer and transfer_prover
kuncar
parents: 61366
diff changeset
   677
      fun transfer_search_tac i =
46af4f577c7e new methods for debugging transfer and transfer_prover
kuncar
parents: 61366
diff changeset
   678
        (SOLVED'
46af4f577c7e new methods for debugging transfer and transfer_prover
kuncar
parents: 61366
diff changeset
   679
          (REPEAT_ALL_NEW (resolve_tac ctxt rules) THEN_ALL_NEW
46af4f577c7e new methods for debugging transfer and transfer_prover
kuncar
parents: 61366
diff changeset
   680
            (DETERM o eq_rules_tac ctxt eq_rules))
46af4f577c7e new methods for debugging transfer and transfer_prover
kuncar
parents: 61366
diff changeset
   681
          ORELSE' end_tac) i
46af4f577c7e new methods for debugging transfer and transfer_prover
kuncar
parents: 61366
diff changeset
   682
     in
46af4f577c7e new methods for debugging transfer and transfer_prover
kuncar
parents: 61366
diff changeset
   683
       (transfer_start_tac equiv ctxt 
46af4f577c7e new methods for debugging transfer and transfer_prover
kuncar
parents: 61366
diff changeset
   684
        THEN_ALL_BUT_FIRST_NEW transfer_search_tac
46af4f577c7e new methods for debugging transfer and transfer_prover
kuncar
parents: 61366
diff changeset
   685
        THEN' transfer_end_tac ctxt) i
46af4f577c7e new methods for debugging transfer and transfer_prover
kuncar
parents: 61366
diff changeset
   686
     end
46af4f577c7e new methods for debugging transfer and transfer_prover
kuncar
parents: 61366
diff changeset
   687
46af4f577c7e new methods for debugging transfer and transfer_prover
kuncar
parents: 61366
diff changeset
   688
  fun transfer_prover_tac ctxt i = 
46af4f577c7e new methods for debugging transfer and transfer_prover
kuncar
parents: 61366
diff changeset
   689
    let
46af4f577c7e new methods for debugging transfer and transfer_prover
kuncar
parents: 61366
diff changeset
   690
      val rules = get_transfer_raw ctxt
46af4f577c7e new methods for debugging transfer and transfer_prover
kuncar
parents: 61366
diff changeset
   691
      val eq_rules = get_relator_eq_raw ctxt
46af4f577c7e new methods for debugging transfer and transfer_prover
kuncar
parents: 61366
diff changeset
   692
  
46af4f577c7e new methods for debugging transfer and transfer_prover
kuncar
parents: 61366
diff changeset
   693
      fun transfer_prover_search_tac i = 
46af4f577c7e new methods for debugging transfer and transfer_prover
kuncar
parents: 61366
diff changeset
   694
        (REPEAT_ALL_NEW (resolve_tac ctxt rules) THEN_ALL_NEW
46af4f577c7e new methods for debugging transfer and transfer_prover
kuncar
parents: 61366
diff changeset
   695
          (DETERM o eq_rules_tac ctxt eq_rules)) i
46af4f577c7e new methods for debugging transfer and transfer_prover
kuncar
parents: 61366
diff changeset
   696
    in
46af4f577c7e new methods for debugging transfer and transfer_prover
kuncar
parents: 61366
diff changeset
   697
      (transfer_prover_start_tac ctxt 
46af4f577c7e new methods for debugging transfer and transfer_prover
kuncar
parents: 61366
diff changeset
   698
       THEN_ALL_BUT_FIRST_NEW transfer_prover_search_tac
46af4f577c7e new methods for debugging transfer and transfer_prover
kuncar
parents: 61366
diff changeset
   699
       THEN' transfer_prover_end_tac ctxt) i
46af4f577c7e new methods for debugging transfer and transfer_prover
kuncar
parents: 61366
diff changeset
   700
    end
46af4f577c7e new methods for debugging transfer and transfer_prover
kuncar
parents: 61366
diff changeset
   701
end;
47789
71a526ee569a implement transfer tactic with more scalable forward proof methods
huffman
parents: 47658
diff changeset
   702
52354
acb4f932dd24 implement 'transferred' attribute for transfer package, with support for monotonicity of !!/==>
huffman
parents: 51996
diff changeset
   703
(** Transfer attribute **)
acb4f932dd24 implement 'transferred' attribute for transfer package, with support for monotonicity of !!/==>
huffman
parents: 51996
diff changeset
   704
acb4f932dd24 implement 'transferred' attribute for transfer package, with support for monotonicity of !!/==>
huffman
parents: 51996
diff changeset
   705
fun transferred ctxt extra_rules thm =
acb4f932dd24 implement 'transferred' attribute for transfer package, with support for monotonicity of !!/==>
huffman
parents: 51996
diff changeset
   706
  let
acb4f932dd24 implement 'transferred' attribute for transfer package, with support for monotonicity of !!/==>
huffman
parents: 51996
diff changeset
   707
    val rules = extra_rules @ get_transfer_raw ctxt
acb4f932dd24 implement 'transferred' attribute for transfer package, with support for monotonicity of !!/==>
huffman
parents: 51996
diff changeset
   708
    val eq_rules = get_relator_eq_raw ctxt
acb4f932dd24 implement 'transferred' attribute for transfer package, with support for monotonicity of !!/==>
huffman
parents: 51996
diff changeset
   709
    val pre_simps = @{thms transfer_forall_eq transfer_implies_eq}
acb4f932dd24 implement 'transferred' attribute for transfer package, with support for monotonicity of !!/==>
huffman
parents: 51996
diff changeset
   710
    val thm1 = Drule.forall_intr_vars thm
60805
4cc49ead6e75 tuned signature;
wenzelm
parents: 60801
diff changeset
   711
    val instT =
4cc49ead6e75 tuned signature;
wenzelm
parents: 60801
diff changeset
   712
      rev (Term.add_tvars (Thm.full_prop_of thm1) [])
4cc49ead6e75 tuned signature;
wenzelm
parents: 60801
diff changeset
   713
      |> map (fn v as ((a, _), S) => (v, Thm.ctyp_of ctxt (TFree (a, S))))
52354
acb4f932dd24 implement 'transferred' attribute for transfer package, with support for monotonicity of !!/==>
huffman
parents: 51996
diff changeset
   714
    val thm2 = thm1
60805
4cc49ead6e75 tuned signature;
wenzelm
parents: 60801
diff changeset
   715
      |> Thm.instantiate (instT, [])
54742
7a86358a3c0b proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
wenzelm
parents: 53649
diff changeset
   716
      |> Raw_Simplifier.rewrite_rule ctxt pre_simps
52354
acb4f932dd24 implement 'transferred' attribute for transfer package, with support for monotonicity of !!/==>
huffman
parents: 51996
diff changeset
   717
    val ctxt' = Variable.declare_names (Thm.full_prop_of thm2) ctxt
70316
c61b7af108a6 misc tuning and clarification, notably wrt. flow of context;
wenzelm
parents: 70159
diff changeset
   718
    val rule = transfer_rule_of_lhs ctxt' (HOLogic.dest_Trueprop (Thm.concl_of thm2))
52354
acb4f932dd24 implement 'transferred' attribute for transfer package, with support for monotonicity of !!/==>
huffman
parents: 51996
diff changeset
   719
  in
70316
c61b7af108a6 misc tuning and clarification, notably wrt. flow of context;
wenzelm
parents: 70159
diff changeset
   720
    Goal.prove_internal ctxt' []
c61b7af108a6 misc tuning and clarification, notably wrt. flow of context;
wenzelm
parents: 70159
diff changeset
   721
      (Thm.cterm_of ctxt' (HOLogic.mk_Trueprop (Var (("P", 0), \<^typ>\<open>bool\<close>))))
c61b7af108a6 misc tuning and clarification, notably wrt. flow of context;
wenzelm
parents: 70159
diff changeset
   722
      (fn _ =>
c61b7af108a6 misc tuning and clarification, notably wrt. flow of context;
wenzelm
parents: 70159
diff changeset
   723
        resolve_tac ctxt' [thm2 RS @{thm transfer_start'}, thm2 RS @{thm transfer_start}] 1 THEN
c61b7af108a6 misc tuning and clarification, notably wrt. flow of context;
wenzelm
parents: 70159
diff changeset
   724
        (resolve_tac ctxt' [rule]
c61b7af108a6 misc tuning and clarification, notably wrt. flow of context;
wenzelm
parents: 70159
diff changeset
   725
          THEN_ALL_NEW
c61b7af108a6 misc tuning and clarification, notably wrt. flow of context;
wenzelm
parents: 70159
diff changeset
   726
            (SOLVED' (REPEAT_ALL_NEW (resolve_tac ctxt' rules)
c61b7af108a6 misc tuning and clarification, notably wrt. flow of context;
wenzelm
parents: 70159
diff changeset
   727
              THEN_ALL_NEW (DETERM o eq_rules_tac ctxt' eq_rules)))) 1
c61b7af108a6 misc tuning and clarification, notably wrt. flow of context;
wenzelm
parents: 70159
diff changeset
   728
          handle TERM (_, ts) =>
c61b7af108a6 misc tuning and clarification, notably wrt. flow of context;
wenzelm
parents: 70159
diff changeset
   729
            raise TERM ("Transfer failed to convert goal to an object-logic formula", ts))
c61b7af108a6 misc tuning and clarification, notably wrt. flow of context;
wenzelm
parents: 70159
diff changeset
   730
    |> Raw_Simplifier.rewrite_rule ctxt' post_simps
c61b7af108a6 misc tuning and clarification, notably wrt. flow of context;
wenzelm
parents: 70159
diff changeset
   731
    |> Simplifier.norm_hhf ctxt'
c61b7af108a6 misc tuning and clarification, notably wrt. flow of context;
wenzelm
parents: 70159
diff changeset
   732
    |> Drule.generalize (map (fst o dest_TFree o Thm.typ_of o snd) instT, [])
c61b7af108a6 misc tuning and clarification, notably wrt. flow of context;
wenzelm
parents: 70159
diff changeset
   733
    |> Drule.zero_var_indexes
52354
acb4f932dd24 implement 'transferred' attribute for transfer package, with support for monotonicity of !!/==>
huffman
parents: 51996
diff changeset
   734
  end
acb4f932dd24 implement 'transferred' attribute for transfer package, with support for monotonicity of !!/==>
huffman
parents: 51996
diff changeset
   735
(*
acb4f932dd24 implement 'transferred' attribute for transfer package, with support for monotonicity of !!/==>
huffman
parents: 51996
diff changeset
   736
    handle THM _ => thm
acb4f932dd24 implement 'transferred' attribute for transfer package, with support for monotonicity of !!/==>
huffman
parents: 51996
diff changeset
   737
*)
acb4f932dd24 implement 'transferred' attribute for transfer package, with support for monotonicity of !!/==>
huffman
parents: 51996
diff changeset
   738
52358
f4c4bcb0d564 implement 'untransferred' attribute, which is like 'transferred' but works in the opposite direction
huffman
parents: 52354
diff changeset
   739
fun untransferred ctxt extra_rules thm =
f4c4bcb0d564 implement 'untransferred' attribute, which is like 'transferred' but works in the opposite direction
huffman
parents: 52354
diff changeset
   740
  let
f4c4bcb0d564 implement 'untransferred' attribute, which is like 'transferred' but works in the opposite direction
huffman
parents: 52354
diff changeset
   741
    val rules = extra_rules @ get_transfer_raw ctxt
f4c4bcb0d564 implement 'untransferred' attribute, which is like 'transferred' but works in the opposite direction
huffman
parents: 52354
diff changeset
   742
    val eq_rules = get_relator_eq_raw ctxt
f4c4bcb0d564 implement 'untransferred' attribute, which is like 'transferred' but works in the opposite direction
huffman
parents: 52354
diff changeset
   743
    val pre_simps = @{thms transfer_forall_eq transfer_implies_eq}
f4c4bcb0d564 implement 'untransferred' attribute, which is like 'transferred' but works in the opposite direction
huffman
parents: 52354
diff changeset
   744
    val thm1 = Drule.forall_intr_vars thm
60805
4cc49ead6e75 tuned signature;
wenzelm
parents: 60801
diff changeset
   745
    val instT =
4cc49ead6e75 tuned signature;
wenzelm
parents: 60801
diff changeset
   746
      rev (Term.add_tvars (Thm.full_prop_of thm1) [])
4cc49ead6e75 tuned signature;
wenzelm
parents: 60801
diff changeset
   747
      |> map (fn v as ((a, _), S) => (v, Thm.ctyp_of ctxt (TFree (a, S))))
52358
f4c4bcb0d564 implement 'untransferred' attribute, which is like 'transferred' but works in the opposite direction
huffman
parents: 52354
diff changeset
   748
    val thm2 = thm1
60805
4cc49ead6e75 tuned signature;
wenzelm
parents: 60801
diff changeset
   749
      |> Thm.instantiate (instT, [])
54742
7a86358a3c0b proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
wenzelm
parents: 53649
diff changeset
   750
      |> Raw_Simplifier.rewrite_rule ctxt pre_simps
52358
f4c4bcb0d564 implement 'untransferred' attribute, which is like 'transferred' but works in the opposite direction
huffman
parents: 52354
diff changeset
   751
    val ctxt' = Variable.declare_names (Thm.full_prop_of thm2) ctxt
f4c4bcb0d564 implement 'untransferred' attribute, which is like 'transferred' but works in the opposite direction
huffman
parents: 52354
diff changeset
   752
    val t = HOLogic.dest_Trueprop (Thm.concl_of thm2)
f4c4bcb0d564 implement 'untransferred' attribute, which is like 'transferred' but works in the opposite direction
huffman
parents: 52354
diff changeset
   753
    val rule = transfer_rule_of_term ctxt' true t
f4c4bcb0d564 implement 'untransferred' attribute, which is like 'transferred' but works in the opposite direction
huffman
parents: 52354
diff changeset
   754
  in
70316
c61b7af108a6 misc tuning and clarification, notably wrt. flow of context;
wenzelm
parents: 70159
diff changeset
   755
    Goal.prove_internal ctxt' []
c61b7af108a6 misc tuning and clarification, notably wrt. flow of context;
wenzelm
parents: 70159
diff changeset
   756
      (Thm.cterm_of ctxt' (HOLogic.mk_Trueprop (Var (("P", 0), \<^typ>\<open>bool\<close>))))
c61b7af108a6 misc tuning and clarification, notably wrt. flow of context;
wenzelm
parents: 70159
diff changeset
   757
      (fn _ =>
c61b7af108a6 misc tuning and clarification, notably wrt. flow of context;
wenzelm
parents: 70159
diff changeset
   758
        resolve_tac ctxt' [thm2 RS @{thm untransfer_start}] 1 THEN
c61b7af108a6 misc tuning and clarification, notably wrt. flow of context;
wenzelm
parents: 70159
diff changeset
   759
        (resolve_tac ctxt' [rule]
c61b7af108a6 misc tuning and clarification, notably wrt. flow of context;
wenzelm
parents: 70159
diff changeset
   760
          THEN_ALL_NEW
c61b7af108a6 misc tuning and clarification, notably wrt. flow of context;
wenzelm
parents: 70159
diff changeset
   761
            (SOLVED' (REPEAT_ALL_NEW (resolve_tac ctxt' rules)
c61b7af108a6 misc tuning and clarification, notably wrt. flow of context;
wenzelm
parents: 70159
diff changeset
   762
              THEN_ALL_NEW (DETERM o eq_rules_tac ctxt' eq_rules)))) 1
c61b7af108a6 misc tuning and clarification, notably wrt. flow of context;
wenzelm
parents: 70159
diff changeset
   763
          handle TERM (_, ts) =>
c61b7af108a6 misc tuning and clarification, notably wrt. flow of context;
wenzelm
parents: 70159
diff changeset
   764
            raise TERM ("Transfer failed to convert goal to an object-logic formula", ts))
c61b7af108a6 misc tuning and clarification, notably wrt. flow of context;
wenzelm
parents: 70159
diff changeset
   765
    |> Raw_Simplifier.rewrite_rule ctxt' post_simps
c61b7af108a6 misc tuning and clarification, notably wrt. flow of context;
wenzelm
parents: 70159
diff changeset
   766
    |> Simplifier.norm_hhf ctxt'
c61b7af108a6 misc tuning and clarification, notably wrt. flow of context;
wenzelm
parents: 70159
diff changeset
   767
    |> Drule.generalize (map (fst o dest_TFree o Thm.typ_of o snd) instT, [])
c61b7af108a6 misc tuning and clarification, notably wrt. flow of context;
wenzelm
parents: 70159
diff changeset
   768
    |> Drule.zero_var_indexes
52358
f4c4bcb0d564 implement 'untransferred' attribute, which is like 'transferred' but works in the opposite direction
huffman
parents: 52354
diff changeset
   769
  end
f4c4bcb0d564 implement 'untransferred' attribute, which is like 'transferred' but works in the opposite direction
huffman
parents: 52354
diff changeset
   770
47789
71a526ee569a implement transfer tactic with more scalable forward proof methods
huffman
parents: 47658
diff changeset
   771
(** Methods and attributes **)
47325
ec6187036495 new transfer proof method
huffman
parents:
diff changeset
   772
47568
98c8b7542b72 add option to transfer method for specifying variables not to generalize over
huffman
parents: 47523
diff changeset
   773
val free = Args.context -- Args.term >> (fn (_, Free v) => v | (ctxt, t) =>
98c8b7542b72 add option to transfer method for specifying variables not to generalize over
huffman
parents: 47523
diff changeset
   774
  error ("Bad free variable: " ^ Syntax.string_of_term ctxt t))
98c8b7542b72 add option to transfer method for specifying variables not to generalize over
huffman
parents: 47523
diff changeset
   775
98c8b7542b72 add option to transfer method for specifying variables not to generalize over
huffman
parents: 47523
diff changeset
   776
val fixing = Scan.optional (Scan.lift (Args.$$$ "fixing" -- Args.colon)
98c8b7542b72 add option to transfer method for specifying variables not to generalize over
huffman
parents: 47523
diff changeset
   777
  |-- Scan.repeat free) []
98c8b7542b72 add option to transfer method for specifying variables not to generalize over
huffman
parents: 47523
diff changeset
   778
61367
46af4f577c7e new methods for debugging transfer and transfer_prover
kuncar
parents: 61366
diff changeset
   779
val reverse_prems = fn _ => PRIMITIVE (fn thm => fold_rev (fn i => Thm.permute_prems i 1) 
46af4f577c7e new methods for debugging transfer and transfer_prover
kuncar
parents: 61366
diff changeset
   780
    (0 upto Thm.nprems_of thm - 1) thm);
46af4f577c7e new methods for debugging transfer and transfer_prover
kuncar
parents: 61366
diff changeset
   781
  
46af4f577c7e new methods for debugging transfer and transfer_prover
kuncar
parents: 61366
diff changeset
   782
fun transfer_start_method equiv : (Proof.context -> Proof.method) context_parser =
46af4f577c7e new methods for debugging transfer and transfer_prover
kuncar
parents: 61366
diff changeset
   783
  fixing >> (fn vs => fn ctxt =>
46af4f577c7e new methods for debugging transfer and transfer_prover
kuncar
parents: 61366
diff changeset
   784
    SIMPLE_METHOD' (gen_frees_tac vs ctxt THEN' transfer_start_tac equiv ctxt
46af4f577c7e new methods for debugging transfer and transfer_prover
kuncar
parents: 61366
diff changeset
   785
    THEN' reverse_prems));
46af4f577c7e new methods for debugging transfer and transfer_prover
kuncar
parents: 61366
diff changeset
   786
53042
aa0322a65bea standardized aliases;
wenzelm
parents: 52884
diff changeset
   787
fun transfer_method equiv : (Proof.context -> Proof.method) context_parser =
47568
98c8b7542b72 add option to transfer method for specifying variables not to generalize over
huffman
parents: 47523
diff changeset
   788
  fixing >> (fn vs => fn ctxt =>
47658
7631f6f7873d enable variant of transfer method that proves an implication instead of an equivalence
huffman
parents: 47635
diff changeset
   789
    SIMPLE_METHOD' (gen_frees_tac vs ctxt THEN' transfer_tac equiv ctxt))
47325
ec6187036495 new transfer proof method
huffman
parents:
diff changeset
   790
61367
46af4f577c7e new methods for debugging transfer and transfer_prover
kuncar
parents: 61366
diff changeset
   791
val transfer_prover_start_method : (Proof.context -> Proof.method) context_parser =
46af4f577c7e new methods for debugging transfer and transfer_prover
kuncar
parents: 61366
diff changeset
   792
  Scan.succeed (fn ctxt => SIMPLE_METHOD' (transfer_prover_start_tac ctxt THEN' reverse_prems))
46af4f577c7e new methods for debugging transfer and transfer_prover
kuncar
parents: 61366
diff changeset
   793
53042
aa0322a65bea standardized aliases;
wenzelm
parents: 52884
diff changeset
   794
val transfer_prover_method : (Proof.context -> Proof.method) context_parser =
47635
ebb79474262c rename 'correspondence' method to 'transfer_prover'
huffman
parents: 47618
diff changeset
   795
  Scan.succeed (fn ctxt => SIMPLE_METHOD' (transfer_prover_tac ctxt))
47325
ec6187036495 new transfer proof method
huffman
parents:
diff changeset
   796
47635
ebb79474262c rename 'correspondence' method to 'transfer_prover'
huffman
parents: 47618
diff changeset
   797
(* Attribute for transfer rules *)
47325
ec6187036495 new transfer proof method
huffman
parents:
diff changeset
   798
58821
11e226e8a095 modernized setup;
wenzelm
parents: 56722
diff changeset
   799
fun prep_rule ctxt =
54742
7a86358a3c0b proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
wenzelm
parents: 53649
diff changeset
   800
  abstract_domains_transfer ctxt o abstract_equalities_transfer ctxt o Conv.fconv_rule prep_conv
47325
ec6187036495 new transfer proof method
huffman
parents:
diff changeset
   801
ec6187036495 new transfer proof method
huffman
parents:
diff changeset
   802
val transfer_add =
58821
11e226e8a095 modernized setup;
wenzelm
parents: 56722
diff changeset
   803
  Thm.declaration_attribute (fn thm => fn ctxt =>
52884
34c47bc771f2 contract equalities in transfer and transfer domain rules when they are registered
kuncar
parents: 52883
diff changeset
   804
    (add_transfer_thm o prep_rule (Context.proof_of ctxt)) thm ctxt)
47325
ec6187036495 new transfer proof method
huffman
parents:
diff changeset
   805
ec6187036495 new transfer proof method
huffman
parents:
diff changeset
   806
val transfer_del =
58821
11e226e8a095 modernized setup;
wenzelm
parents: 56722
diff changeset
   807
  Thm.declaration_attribute (fn thm => fn ctxt =>
52884
34c47bc771f2 contract equalities in transfer and transfer domain rules when they are registered
kuncar
parents: 52883
diff changeset
   808
    (del_transfer_thm o prep_rule (Context.proof_of ctxt)) thm ctxt)
47325
ec6187036495 new transfer proof method
huffman
parents:
diff changeset
   809
ec6187036495 new transfer proof method
huffman
parents:
diff changeset
   810
val transfer_attribute =
ec6187036495 new transfer proof method
huffman
parents:
diff changeset
   811
  Attrib.add_del transfer_add transfer_del
ec6187036495 new transfer proof method
huffman
parents:
diff changeset
   812
51956
a4d81cdebf8b better support for domains in Lifting/Transfer = replace Domainp T by the actual invariant in a transferred goal
kuncar
parents: 51955
diff changeset
   813
(* Attributes for transfer domain rules *)
a4d81cdebf8b better support for domains in Lifting/Transfer = replace Domainp T by the actual invariant in a transferred goal
kuncar
parents: 51955
diff changeset
   814
a4d81cdebf8b better support for domains in Lifting/Transfer = replace Domainp T by the actual invariant in a transferred goal
kuncar
parents: 51955
diff changeset
   815
val transfer_domain_add = Thm.declaration_attribute add_transfer_domain_thm
a4d81cdebf8b better support for domains in Lifting/Transfer = replace Domainp T by the actual invariant in a transferred goal
kuncar
parents: 51955
diff changeset
   816
a4d81cdebf8b better support for domains in Lifting/Transfer = replace Domainp T by the actual invariant in a transferred goal
kuncar
parents: 51955
diff changeset
   817
val transfer_domain_del = Thm.declaration_attribute del_transfer_domain_thm
a4d81cdebf8b better support for domains in Lifting/Transfer = replace Domainp T by the actual invariant in a transferred goal
kuncar
parents: 51955
diff changeset
   818
a4d81cdebf8b better support for domains in Lifting/Transfer = replace Domainp T by the actual invariant in a transferred goal
kuncar
parents: 51955
diff changeset
   819
val transfer_domain_attribute =
a4d81cdebf8b better support for domains in Lifting/Transfer = replace Domainp T by the actual invariant in a transferred goal
kuncar
parents: 51955
diff changeset
   820
  Attrib.add_del transfer_domain_add transfer_domain_del
a4d81cdebf8b better support for domains in Lifting/Transfer = replace Domainp T by the actual invariant in a transferred goal
kuncar
parents: 51955
diff changeset
   821
52354
acb4f932dd24 implement 'transferred' attribute for transfer package, with support for monotonicity of !!/==>
huffman
parents: 51996
diff changeset
   822
(* Attributes for transferred rules *)
acb4f932dd24 implement 'transferred' attribute for transfer package, with support for monotonicity of !!/==>
huffman
parents: 51996
diff changeset
   823
61853
fb7756087101 rule_attribute and declaration_attribute implicitly support abstract closure, but mixed_attribute implementations need to be aware of Thm.is_free_dummy;
wenzelm
parents: 61367
diff changeset
   824
fun transferred_attribute thms =
fb7756087101 rule_attribute and declaration_attribute implicitly support abstract closure, but mixed_attribute implementations need to be aware of Thm.is_free_dummy;
wenzelm
parents: 61367
diff changeset
   825
  Thm.rule_attribute thms (fn context => transferred (Context.proof_of context) thms)
52354
acb4f932dd24 implement 'transferred' attribute for transfer package, with support for monotonicity of !!/==>
huffman
parents: 51996
diff changeset
   826
61853
fb7756087101 rule_attribute and declaration_attribute implicitly support abstract closure, but mixed_attribute implementations need to be aware of Thm.is_free_dummy;
wenzelm
parents: 61367
diff changeset
   827
fun untransferred_attribute thms =
fb7756087101 rule_attribute and declaration_attribute implicitly support abstract closure, but mixed_attribute implementations need to be aware of Thm.is_free_dummy;
wenzelm
parents: 61367
diff changeset
   828
  Thm.rule_attribute thms (fn context => untransferred (Context.proof_of context) thms)
52358
f4c4bcb0d564 implement 'untransferred' attribute, which is like 'transferred' but works in the opposite direction
huffman
parents: 52354
diff changeset
   829
52354
acb4f932dd24 implement 'transferred' attribute for transfer package, with support for monotonicity of !!/==>
huffman
parents: 51996
diff changeset
   830
val transferred_attribute_parser =
acb4f932dd24 implement 'transferred' attribute for transfer package, with support for monotonicity of !!/==>
huffman
parents: 51996
diff changeset
   831
  Attrib.thms >> transferred_attribute
acb4f932dd24 implement 'transferred' attribute for transfer package, with support for monotonicity of !!/==>
huffman
parents: 51996
diff changeset
   832
52358
f4c4bcb0d564 implement 'untransferred' attribute, which is like 'transferred' but works in the opposite direction
huffman
parents: 52354
diff changeset
   833
val untransferred_attribute_parser =
f4c4bcb0d564 implement 'untransferred' attribute, which is like 'transferred' but works in the opposite direction
huffman
parents: 52354
diff changeset
   834
  Attrib.thms >> untransferred_attribute
f4c4bcb0d564 implement 'untransferred' attribute, which is like 'transferred' but works in the opposite direction
huffman
parents: 52354
diff changeset
   835
60220
530112e8c6ba tuned; store pred_simps
kuncar
parents: 60216
diff changeset
   836
fun morph_pred_data phi = 
530112e8c6ba tuned; store pred_simps
kuncar
parents: 60216
diff changeset
   837
  let
530112e8c6ba tuned; store pred_simps
kuncar
parents: 60216
diff changeset
   838
    val morph_thm = Morphism.thm phi
530112e8c6ba tuned; store pred_simps
kuncar
parents: 60216
diff changeset
   839
  in
530112e8c6ba tuned; store pred_simps
kuncar
parents: 60216
diff changeset
   840
    map_pred_data' morph_thm morph_thm (map morph_thm)
530112e8c6ba tuned; store pred_simps
kuncar
parents: 60216
diff changeset
   841
  end
56524
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents: 56520
diff changeset
   842
70316
c61b7af108a6 misc tuning and clarification, notably wrt. flow of context;
wenzelm
parents: 70159
diff changeset
   843
fun lookup_pred_data ctxt type_name =
c61b7af108a6 misc tuning and clarification, notably wrt. flow of context;
wenzelm
parents: 70159
diff changeset
   844
  Symtab.lookup (get_pred_data ctxt) type_name
67664
ad2b3e330c27 tuned signature;
wenzelm
parents: 67399
diff changeset
   845
  |> Option.map (morph_pred_data (Morphism.transfer_morphism' ctxt))
56524
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents: 56520
diff changeset
   846
58821
11e226e8a095 modernized setup;
wenzelm
parents: 56722
diff changeset
   847
fun update_pred_data type_name qinfo ctxt =
56524
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents: 56520
diff changeset
   848
  Data.map (map_pred_data (Symtab.update (type_name, qinfo))) ctxt
f4ba736040fa setup for Transfer and Lifting from BNF; tuned thm names
kuncar
parents: 56520
diff changeset
   849
47325
ec6187036495 new transfer proof method
huffman
parents:
diff changeset
   850
(* Theory setup *)
ec6187036495 new transfer proof method
huffman
parents:
diff changeset
   851
58821
11e226e8a095 modernized setup;
wenzelm
parents: 56722
diff changeset
   852
val _ =
11e226e8a095 modernized setup;
wenzelm
parents: 56722
diff changeset
   853
  Theory.setup
11e226e8a095 modernized setup;
wenzelm
parents: 56722
diff changeset
   854
    let
69593
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 67664
diff changeset
   855
      val name = \<^binding>\<open>relator_eq\<close>
58821
11e226e8a095 modernized setup;
wenzelm
parents: 56722
diff changeset
   856
      fun add_thm thm context = context
11e226e8a095 modernized setup;
wenzelm
parents: 56722
diff changeset
   857
        |> Data.map (map_relator_eq (Item_Net.update thm))
11e226e8a095 modernized setup;
wenzelm
parents: 56722
diff changeset
   858
        |> Data.map (map_relator_eq_raw
11e226e8a095 modernized setup;
wenzelm
parents: 56722
diff changeset
   859
            (Item_Net.update (abstract_equalities_relator_eq (Context.proof_of context) thm)))
11e226e8a095 modernized setup;
wenzelm
parents: 56722
diff changeset
   860
      fun del_thm thm context = context
11e226e8a095 modernized setup;
wenzelm
parents: 56722
diff changeset
   861
        |> Data.map (map_relator_eq (Item_Net.remove thm))
11e226e8a095 modernized setup;
wenzelm
parents: 56722
diff changeset
   862
        |> Data.map (map_relator_eq_raw
11e226e8a095 modernized setup;
wenzelm
parents: 56722
diff changeset
   863
            (Item_Net.remove (abstract_equalities_relator_eq (Context.proof_of context) thm)))
11e226e8a095 modernized setup;
wenzelm
parents: 56722
diff changeset
   864
      val add = Thm.declaration_attribute add_thm
11e226e8a095 modernized setup;
wenzelm
parents: 56722
diff changeset
   865
      val del = Thm.declaration_attribute del_thm
11e226e8a095 modernized setup;
wenzelm
parents: 56722
diff changeset
   866
      val text = "declaration of relator equality rule (used by transfer method)"
11e226e8a095 modernized setup;
wenzelm
parents: 56722
diff changeset
   867
      val content = Item_Net.content o #relator_eq o Data.get
11e226e8a095 modernized setup;
wenzelm
parents: 56722
diff changeset
   868
    in
11e226e8a095 modernized setup;
wenzelm
parents: 56722
diff changeset
   869
      Attrib.setup name (Attrib.add_del add del) text
11e226e8a095 modernized setup;
wenzelm
parents: 56722
diff changeset
   870
      #> Global_Theory.add_thms_dynamic (name, content)
11e226e8a095 modernized setup;
wenzelm
parents: 56722
diff changeset
   871
    end
49975
faf4afed009f transfer package: more flexible handling of equality relations using is_equality predicate
huffman
parents: 49625
diff changeset
   872
58821
11e226e8a095 modernized setup;
wenzelm
parents: 56722
diff changeset
   873
val _ =
11e226e8a095 modernized setup;
wenzelm
parents: 56722
diff changeset
   874
  Theory.setup
11e226e8a095 modernized setup;
wenzelm
parents: 56722
diff changeset
   875
    let
69593
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 67664
diff changeset
   876
      val name = \<^binding>\<open>relator_domain\<close>
58821
11e226e8a095 modernized setup;
wenzelm
parents: 56722
diff changeset
   877
      fun add_thm thm context =
11e226e8a095 modernized setup;
wenzelm
parents: 56722
diff changeset
   878
        let
11e226e8a095 modernized setup;
wenzelm
parents: 56722
diff changeset
   879
          val thm = abstract_domains_relator_domain (Context.proof_of context) thm
11e226e8a095 modernized setup;
wenzelm
parents: 56722
diff changeset
   880
        in
11e226e8a095 modernized setup;
wenzelm
parents: 56722
diff changeset
   881
          context |> Data.map (map_relator_domain (Item_Net.update thm)) |> add_transfer_domain_thm thm
11e226e8a095 modernized setup;
wenzelm
parents: 56722
diff changeset
   882
        end
11e226e8a095 modernized setup;
wenzelm
parents: 56722
diff changeset
   883
      fun del_thm thm context =
11e226e8a095 modernized setup;
wenzelm
parents: 56722
diff changeset
   884
        let
11e226e8a095 modernized setup;
wenzelm
parents: 56722
diff changeset
   885
          val thm = abstract_domains_relator_domain (Context.proof_of context) thm
11e226e8a095 modernized setup;
wenzelm
parents: 56722
diff changeset
   886
        in
11e226e8a095 modernized setup;
wenzelm
parents: 56722
diff changeset
   887
          context |> Data.map (map_relator_domain (Item_Net.remove thm)) |> del_transfer_domain_thm thm
11e226e8a095 modernized setup;
wenzelm
parents: 56722
diff changeset
   888
        end
11e226e8a095 modernized setup;
wenzelm
parents: 56722
diff changeset
   889
      val add = Thm.declaration_attribute add_thm
11e226e8a095 modernized setup;
wenzelm
parents: 56722
diff changeset
   890
      val del = Thm.declaration_attribute del_thm
11e226e8a095 modernized setup;
wenzelm
parents: 56722
diff changeset
   891
      val text = "declaration of relator domain rule (used by transfer method)"
11e226e8a095 modernized setup;
wenzelm
parents: 56722
diff changeset
   892
      val content = Item_Net.content o #relator_domain o Data.get
11e226e8a095 modernized setup;
wenzelm
parents: 56722
diff changeset
   893
    in
11e226e8a095 modernized setup;
wenzelm
parents: 56722
diff changeset
   894
      Attrib.setup name (Attrib.add_del add del) text
11e226e8a095 modernized setup;
wenzelm
parents: 56722
diff changeset
   895
      #> Global_Theory.add_thms_dynamic (name, content)
11e226e8a095 modernized setup;
wenzelm
parents: 56722
diff changeset
   896
    end
51956
a4d81cdebf8b better support for domains in Lifting/Transfer = replace Domainp T by the actual invariant in a transferred goal
kuncar
parents: 51955
diff changeset
   897
58821
11e226e8a095 modernized setup;
wenzelm
parents: 56722
diff changeset
   898
val _ =
11e226e8a095 modernized setup;
wenzelm
parents: 56722
diff changeset
   899
  Theory.setup
69593
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 67664
diff changeset
   900
    (Attrib.setup \<^binding>\<open>transfer_rule\<close> transfer_attribute
58821
11e226e8a095 modernized setup;
wenzelm
parents: 56722
diff changeset
   901
       "transfer rule for transfer method"
11e226e8a095 modernized setup;
wenzelm
parents: 56722
diff changeset
   902
    #> Global_Theory.add_thms_dynamic
69593
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 67664
diff changeset
   903
       (\<^binding>\<open>transfer_raw\<close>, Item_Net.content o #transfer_raw o Data.get)
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 67664
diff changeset
   904
    #> Attrib.setup \<^binding>\<open>transfer_domain_rule\<close> transfer_domain_attribute
58821
11e226e8a095 modernized setup;
wenzelm
parents: 56722
diff changeset
   905
       "transfer domain rule for transfer method"
69593
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 67664
diff changeset
   906
    #> Attrib.setup \<^binding>\<open>transferred\<close> transferred_attribute_parser
58821
11e226e8a095 modernized setup;
wenzelm
parents: 56722
diff changeset
   907
       "raw theorem transferred to abstract theorem using transfer rules"
69593
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 67664
diff changeset
   908
    #> Attrib.setup \<^binding>\<open>untransferred\<close> untransferred_attribute_parser
58821
11e226e8a095 modernized setup;
wenzelm
parents: 56722
diff changeset
   909
       "abstract theorem transferred to raw theorem using transfer rules"
11e226e8a095 modernized setup;
wenzelm
parents: 56722
diff changeset
   910
    #> Global_Theory.add_thms_dynamic
69593
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 67664
diff changeset
   911
       (\<^binding>\<open>relator_eq_raw\<close>, Item_Net.content o #relator_eq_raw o Data.get)
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 67664
diff changeset
   912
    #> Method.setup \<^binding>\<open>transfer_start\<close> (transfer_start_method true)
61367
46af4f577c7e new methods for debugging transfer and transfer_prover
kuncar
parents: 61366
diff changeset
   913
       "firtst step in the transfer algorithm (for debugging transfer)"
69593
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 67664
diff changeset
   914
    #> Method.setup \<^binding>\<open>transfer_start'\<close> (transfer_start_method false)
61367
46af4f577c7e new methods for debugging transfer and transfer_prover
kuncar
parents: 61366
diff changeset
   915
       "firtst step in the transfer algorithm (for debugging transfer)"
69593
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 67664
diff changeset
   916
    #> Method.setup \<^binding>\<open>transfer_prover_start\<close> transfer_prover_start_method
61367
46af4f577c7e new methods for debugging transfer and transfer_prover
kuncar
parents: 61366
diff changeset
   917
       "firtst step in the transfer_prover algorithm (for debugging transfer_prover)"
69593
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 67664
diff changeset
   918
    #> Method.setup \<^binding>\<open>transfer_step\<close>
61367
46af4f577c7e new methods for debugging transfer and transfer_prover
kuncar
parents: 61366
diff changeset
   919
       (Scan.succeed (fn ctxt => SIMPLE_METHOD' (transfer_step_tac ctxt)))
46af4f577c7e new methods for debugging transfer and transfer_prover
kuncar
parents: 61366
diff changeset
   920
       "step in the search for transfer rules (for debugging transfer and transfer_prover)"
69593
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 67664
diff changeset
   921
    #> Method.setup \<^binding>\<open>transfer_end\<close>
61367
46af4f577c7e new methods for debugging transfer and transfer_prover
kuncar
parents: 61366
diff changeset
   922
       (Scan.succeed (fn ctxt => SIMPLE_METHOD' (transfer_end_tac ctxt)))
46af4f577c7e new methods for debugging transfer and transfer_prover
kuncar
parents: 61366
diff changeset
   923
       "last step in the transfer algorithm (for debugging transfer)"
69593
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 67664
diff changeset
   924
    #> Method.setup \<^binding>\<open>transfer_prover_end\<close>
61367
46af4f577c7e new methods for debugging transfer and transfer_prover
kuncar
parents: 61366
diff changeset
   925
       (Scan.succeed (fn ctxt => SIMPLE_METHOD' (transfer_prover_end_tac ctxt)))
46af4f577c7e new methods for debugging transfer and transfer_prover
kuncar
parents: 61366
diff changeset
   926
       "last step in the transfer_prover algorithm (for debugging transfer_prover)"
69593
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 67664
diff changeset
   927
    #> Method.setup \<^binding>\<open>transfer\<close> (transfer_method true)
58821
11e226e8a095 modernized setup;
wenzelm
parents: 56722
diff changeset
   928
       "generic theorem transfer method"
69593
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 67664
diff changeset
   929
    #> Method.setup \<^binding>\<open>transfer'\<close> (transfer_method false)
58821
11e226e8a095 modernized setup;
wenzelm
parents: 56722
diff changeset
   930
       "generic theorem transfer method"
69593
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 67664
diff changeset
   931
    #> Method.setup \<^binding>\<open>transfer_prover\<close> transfer_prover_method
58821
11e226e8a095 modernized setup;
wenzelm
parents: 56722
diff changeset
   932
       "for proving transfer rules")
47325
ec6187036495 new transfer proof method
huffman
parents:
diff changeset
   933
end