Nominal/Ex/LamTest.thy
author Christian Urban <urbanc@in.tum.de>
Sat, 15 Jan 2011 21:16:15 +0000
changeset 2661 16896fd8eff5
parent 2660 3342a2d13d95
child 2662 7c5bca978886
permissions -rw-r--r--
subst also works now
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
2649
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     1
theory LamTest
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     2
imports "../Nominal2" 
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     3
begin
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     4
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     5
atom_decl name
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     6
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     7
nominal_datatype lam =
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     8
  Var "name"
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     9
| App "lam" "lam"
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    10
| Lam x::"name" l::"lam"  bind x in l
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    11
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    12
definition
2653
d0f774d06e6e added eqvt_at premises in function definition - however not proved at the moment
Christian Urban <urbanc@in.tum.de>
parents: 2652
diff changeset
    13
 "eqvt_at f x \<equiv> \<forall>p. p \<bullet> (f x) = f (p \<bullet> x)"
d0f774d06e6e added eqvt_at premises in function definition - however not proved at the moment
Christian Urban <urbanc@in.tum.de>
parents: 2652
diff changeset
    14
2656
33a6b690fb53 added a property about finite support in the presense of eqvt_at
Christian Urban <urbanc@in.tum.de>
parents: 2655
diff changeset
    15
lemma supp_eqvt_at:
33a6b690fb53 added a property about finite support in the presense of eqvt_at
Christian Urban <urbanc@in.tum.de>
parents: 2655
diff changeset
    16
  assumes asm: "eqvt_at f x"
33a6b690fb53 added a property about finite support in the presense of eqvt_at
Christian Urban <urbanc@in.tum.de>
parents: 2655
diff changeset
    17
  and     fin: "finite (supp x)"
33a6b690fb53 added a property about finite support in the presense of eqvt_at
Christian Urban <urbanc@in.tum.de>
parents: 2655
diff changeset
    18
  shows "supp (f x) \<subseteq> supp x"
33a6b690fb53 added a property about finite support in the presense of eqvt_at
Christian Urban <urbanc@in.tum.de>
parents: 2655
diff changeset
    19
apply(rule supp_is_subset)
33a6b690fb53 added a property about finite support in the presense of eqvt_at
Christian Urban <urbanc@in.tum.de>
parents: 2655
diff changeset
    20
unfolding supports_def
33a6b690fb53 added a property about finite support in the presense of eqvt_at
Christian Urban <urbanc@in.tum.de>
parents: 2655
diff changeset
    21
unfolding fresh_def[symmetric]
33a6b690fb53 added a property about finite support in the presense of eqvt_at
Christian Urban <urbanc@in.tum.de>
parents: 2655
diff changeset
    22
using asm
33a6b690fb53 added a property about finite support in the presense of eqvt_at
Christian Urban <urbanc@in.tum.de>
parents: 2655
diff changeset
    23
apply(simp add: eqvt_at_def)
33a6b690fb53 added a property about finite support in the presense of eqvt_at
Christian Urban <urbanc@in.tum.de>
parents: 2655
diff changeset
    24
apply(simp add: swap_fresh_fresh)
33a6b690fb53 added a property about finite support in the presense of eqvt_at
Christian Urban <urbanc@in.tum.de>
parents: 2655
diff changeset
    25
apply(rule fin)
33a6b690fb53 added a property about finite support in the presense of eqvt_at
Christian Urban <urbanc@in.tum.de>
parents: 2655
diff changeset
    26
done
33a6b690fb53 added a property about finite support in the presense of eqvt_at
Christian Urban <urbanc@in.tum.de>
parents: 2655
diff changeset
    27
2657
1ea9c059fc0f a few lemmas about freshness for at and at_base
Christian Urban <urbanc@in.tum.de>
parents: 2656
diff changeset
    28
lemma finite_supp_eqvt_at:
1ea9c059fc0f a few lemmas about freshness for at and at_base
Christian Urban <urbanc@in.tum.de>
parents: 2656
diff changeset
    29
  assumes asm: "eqvt_at f x"
1ea9c059fc0f a few lemmas about freshness for at and at_base
Christian Urban <urbanc@in.tum.de>
parents: 2656
diff changeset
    30
  and     fin: "finite (supp x)"
1ea9c059fc0f a few lemmas about freshness for at and at_base
Christian Urban <urbanc@in.tum.de>
parents: 2656
diff changeset
    31
  shows "finite (supp (f x))"
1ea9c059fc0f a few lemmas about freshness for at and at_base
Christian Urban <urbanc@in.tum.de>
parents: 2656
diff changeset
    32
apply(rule finite_subset)
1ea9c059fc0f a few lemmas about freshness for at and at_base
Christian Urban <urbanc@in.tum.de>
parents: 2656
diff changeset
    33
apply(rule supp_eqvt_at[OF asm fin])
1ea9c059fc0f a few lemmas about freshness for at and at_base
Christian Urban <urbanc@in.tum.de>
parents: 2656
diff changeset
    34
apply(rule fin)
1ea9c059fc0f a few lemmas about freshness for at and at_base
Christian Urban <urbanc@in.tum.de>
parents: 2656
diff changeset
    35
done
2656
33a6b690fb53 added a property about finite support in the presense of eqvt_at
Christian Urban <urbanc@in.tum.de>
parents: 2655
diff changeset
    36
2660
3342a2d13d95 nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
parents: 2657
diff changeset
    37
lemma fresh_eqvt_at:
3342a2d13d95 nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
parents: 2657
diff changeset
    38
  assumes asm: "eqvt_at f x"
3342a2d13d95 nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
parents: 2657
diff changeset
    39
  and     fin: "finite (supp x)"
3342a2d13d95 nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
parents: 2657
diff changeset
    40
  and     fresh: "as \<sharp>* x"
3342a2d13d95 nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
parents: 2657
diff changeset
    41
  shows "as \<sharp>* f x"
3342a2d13d95 nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
parents: 2657
diff changeset
    42
using fresh
3342a2d13d95 nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
parents: 2657
diff changeset
    43
unfolding fresh_star_def
3342a2d13d95 nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
parents: 2657
diff changeset
    44
unfolding fresh_def
3342a2d13d95 nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
parents: 2657
diff changeset
    45
using supp_eqvt_at[OF asm fin]
3342a2d13d95 nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
parents: 2657
diff changeset
    46
by auto
3342a2d13d95 nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
parents: 2657
diff changeset
    47
2653
d0f774d06e6e added eqvt_at premises in function definition - however not proved at the moment
Christian Urban <urbanc@in.tum.de>
parents: 2652
diff changeset
    48
definition
d0f774d06e6e added eqvt_at premises in function definition - however not proved at the moment
Christian Urban <urbanc@in.tum.de>
parents: 2652
diff changeset
    49
 "eqvt f \<equiv> \<forall>p. p \<bullet> f = f"
2649
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    50
2650
e5fa8de0e4bd derived equivariance for the function graph and function relation
Christian Urban <urbanc@in.tum.de>
parents: 2649
diff changeset
    51
lemma eqvtI:
2653
d0f774d06e6e added eqvt_at premises in function definition - however not proved at the moment
Christian Urban <urbanc@in.tum.de>
parents: 2652
diff changeset
    52
  "(\<And>p. p \<bullet> f \<equiv> f) \<Longrightarrow> eqvt f"
2650
e5fa8de0e4bd derived equivariance for the function graph and function relation
Christian Urban <urbanc@in.tum.de>
parents: 2649
diff changeset
    53
apply(auto simp add: eqvt_def)
e5fa8de0e4bd derived equivariance for the function graph and function relation
Christian Urban <urbanc@in.tum.de>
parents: 2649
diff changeset
    54
done
2649
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    55
2651
4aa72a88b2c1 equivariance of THE_default under the uniqueness assumption
Christian Urban <urbanc@in.tum.de>
parents: 2650
diff changeset
    56
lemma the_default_eqvt:
4aa72a88b2c1 equivariance of THE_default under the uniqueness assumption
Christian Urban <urbanc@in.tum.de>
parents: 2650
diff changeset
    57
  assumes unique: "\<exists>!x. P x"
2652
e9a2770660ef added one further lemma about equivariance of THE_default
Christian Urban <urbanc@in.tum.de>
parents: 2651
diff changeset
    58
  shows "(p \<bullet> (THE_default d P)) = (THE_default d (p \<bullet> P))"
2651
4aa72a88b2c1 equivariance of THE_default under the uniqueness assumption
Christian Urban <urbanc@in.tum.de>
parents: 2650
diff changeset
    59
  apply(rule THE_default1_equality [symmetric])
4aa72a88b2c1 equivariance of THE_default under the uniqueness assumption
Christian Urban <urbanc@in.tum.de>
parents: 2650
diff changeset
    60
  apply(rule_tac p="-p" in permute_boolE)
4aa72a88b2c1 equivariance of THE_default under the uniqueness assumption
Christian Urban <urbanc@in.tum.de>
parents: 2650
diff changeset
    61
  apply(perm_simp add: permute_minus_cancel)
4aa72a88b2c1 equivariance of THE_default under the uniqueness assumption
Christian Urban <urbanc@in.tum.de>
parents: 2650
diff changeset
    62
  apply(rule unique)
4aa72a88b2c1 equivariance of THE_default under the uniqueness assumption
Christian Urban <urbanc@in.tum.de>
parents: 2650
diff changeset
    63
  apply(rule_tac p="-p" in permute_boolE)
4aa72a88b2c1 equivariance of THE_default under the uniqueness assumption
Christian Urban <urbanc@in.tum.de>
parents: 2650
diff changeset
    64
  apply(perm_simp add: permute_minus_cancel)
4aa72a88b2c1 equivariance of THE_default under the uniqueness assumption
Christian Urban <urbanc@in.tum.de>
parents: 2650
diff changeset
    65
  apply(rule THE_defaultI'[OF unique])
4aa72a88b2c1 equivariance of THE_default under the uniqueness assumption
Christian Urban <urbanc@in.tum.de>
parents: 2650
diff changeset
    66
  done
4aa72a88b2c1 equivariance of THE_default under the uniqueness assumption
Christian Urban <urbanc@in.tum.de>
parents: 2650
diff changeset
    67
2652
e9a2770660ef added one further lemma about equivariance of THE_default
Christian Urban <urbanc@in.tum.de>
parents: 2651
diff changeset
    68
lemma fundef_ex1_eqvt:
e9a2770660ef added one further lemma about equivariance of THE_default
Christian Urban <urbanc@in.tum.de>
parents: 2651
diff changeset
    69
  fixes x::"'a::pt"
e9a2770660ef added one further lemma about equivariance of THE_default
Christian Urban <urbanc@in.tum.de>
parents: 2651
diff changeset
    70
  assumes f_def: "f == (\<lambda>x::'a. THE_default d (G x))"
2655
1c3ad1256f16 instantiated fundef_ex1_eqvt_at theorem with the indction hypothesis
Christian Urban <urbanc@in.tum.de>
parents: 2654
diff changeset
    71
  assumes eqvt: "eqvt G"
2652
e9a2770660ef added one further lemma about equivariance of THE_default
Christian Urban <urbanc@in.tum.de>
parents: 2651
diff changeset
    72
  assumes ex1: "\<exists>!y. G x y"
e9a2770660ef added one further lemma about equivariance of THE_default
Christian Urban <urbanc@in.tum.de>
parents: 2651
diff changeset
    73
  shows "(p \<bullet> (f x)) = f (p \<bullet> x)"
e9a2770660ef added one further lemma about equivariance of THE_default
Christian Urban <urbanc@in.tum.de>
parents: 2651
diff changeset
    74
  apply (simp only: f_def)
e9a2770660ef added one further lemma about equivariance of THE_default
Christian Urban <urbanc@in.tum.de>
parents: 2651
diff changeset
    75
  apply(subst the_default_eqvt)
e9a2770660ef added one further lemma about equivariance of THE_default
Christian Urban <urbanc@in.tum.de>
parents: 2651
diff changeset
    76
  apply (rule ex1)
e9a2770660ef added one further lemma about equivariance of THE_default
Christian Urban <urbanc@in.tum.de>
parents: 2651
diff changeset
    77
  apply(perm_simp)
e9a2770660ef added one further lemma about equivariance of THE_default
Christian Urban <urbanc@in.tum.de>
parents: 2651
diff changeset
    78
  using eqvt
e9a2770660ef added one further lemma about equivariance of THE_default
Christian Urban <urbanc@in.tum.de>
parents: 2651
diff changeset
    79
  apply(simp add: eqvt_def)
e9a2770660ef added one further lemma about equivariance of THE_default
Christian Urban <urbanc@in.tum.de>
parents: 2651
diff changeset
    80
  done
e9a2770660ef added one further lemma about equivariance of THE_default
Christian Urban <urbanc@in.tum.de>
parents: 2651
diff changeset
    81
2655
1c3ad1256f16 instantiated fundef_ex1_eqvt_at theorem with the indction hypothesis
Christian Urban <urbanc@in.tum.de>
parents: 2654
diff changeset
    82
lemma fundef_ex1_eqvt_at:
1c3ad1256f16 instantiated fundef_ex1_eqvt_at theorem with the indction hypothesis
Christian Urban <urbanc@in.tum.de>
parents: 2654
diff changeset
    83
  fixes x::"'a::pt"
1c3ad1256f16 instantiated fundef_ex1_eqvt_at theorem with the indction hypothesis
Christian Urban <urbanc@in.tum.de>
parents: 2654
diff changeset
    84
  assumes f_def: "f == (\<lambda>x::'a. THE_default d (G x))"
1c3ad1256f16 instantiated fundef_ex1_eqvt_at theorem with the indction hypothesis
Christian Urban <urbanc@in.tum.de>
parents: 2654
diff changeset
    85
  assumes eqvt: "eqvt G"
1c3ad1256f16 instantiated fundef_ex1_eqvt_at theorem with the indction hypothesis
Christian Urban <urbanc@in.tum.de>
parents: 2654
diff changeset
    86
  assumes ex1: "\<exists>!y. G x y"
1c3ad1256f16 instantiated fundef_ex1_eqvt_at theorem with the indction hypothesis
Christian Urban <urbanc@in.tum.de>
parents: 2654
diff changeset
    87
  shows "eqvt_at f x"
1c3ad1256f16 instantiated fundef_ex1_eqvt_at theorem with the indction hypothesis
Christian Urban <urbanc@in.tum.de>
parents: 2654
diff changeset
    88
  unfolding eqvt_at_def
1c3ad1256f16 instantiated fundef_ex1_eqvt_at theorem with the indction hypothesis
Christian Urban <urbanc@in.tum.de>
parents: 2654
diff changeset
    89
  using assms
1c3ad1256f16 instantiated fundef_ex1_eqvt_at theorem with the indction hypothesis
Christian Urban <urbanc@in.tum.de>
parents: 2654
diff changeset
    90
  apply(auto intro: fundef_ex1_eqvt)
1c3ad1256f16 instantiated fundef_ex1_eqvt_at theorem with the indction hypothesis
Christian Urban <urbanc@in.tum.de>
parents: 2654
diff changeset
    91
  done
2651
4aa72a88b2c1 equivariance of THE_default under the uniqueness assumption
Christian Urban <urbanc@in.tum.de>
parents: 2650
diff changeset
    92
2649
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    93
ML {*
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    94
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    95
val trace = Unsynchronized.ref false
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    96
fun trace_msg msg = if ! trace then tracing (msg ()) else ()
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    97
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    98
val boolT = HOLogic.boolT
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    99
val mk_eq = HOLogic.mk_eq
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   100
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   101
open Function_Lib
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   102
open Function_Common
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   103
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   104
datatype globals = Globals of
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   105
 {fvar: term,
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   106
  domT: typ,
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   107
  ranT: typ,
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   108
  h: term,
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   109
  y: term,
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   110
  x: term,
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   111
  z: term,
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   112
  a: term,
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   113
  P: term,
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   114
  D: term,
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   115
  Pbool:term}
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   116
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   117
datatype rec_call_info = RCInfo of
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   118
 {RIvs: (string * typ) list,  (* Call context: fixes and assumes *)
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   119
  CCas: thm list,
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   120
  rcarg: term,                 (* The recursive argument *)
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   121
  llRI: thm,
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   122
  h_assum: term}
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   123
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   124
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   125
datatype clause_context = ClauseContext of
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   126
 {ctxt : Proof.context,
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   127
  qs : term list,
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   128
  gs : term list,
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   129
  lhs: term,
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   130
  rhs: term,
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   131
  cqs: cterm list,
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   132
  ags: thm list,
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   133
  case_hyp : thm}
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   134
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   135
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   136
fun transfer_clause_ctx thy (ClauseContext { ctxt, qs, gs, lhs, rhs, cqs, ags, case_hyp }) =
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   137
  ClauseContext { ctxt = ProofContext.transfer thy ctxt,
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   138
    qs = qs, gs = gs, lhs = lhs, rhs = rhs, cqs = cqs, ags = ags, case_hyp = case_hyp }
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   139
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   140
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   141
datatype clause_info = ClauseInfo of
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   142
 {no: int,
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   143
  qglr : ((string * typ) list * term list * term * term),
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   144
  cdata : clause_context,
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   145
  tree: Function_Ctx_Tree.ctx_tree,
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   146
  lGI: thm,
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   147
  RCs: rec_call_info list}
2651
4aa72a88b2c1 equivariance of THE_default under the uniqueness assumption
Christian Urban <urbanc@in.tum.de>
parents: 2650
diff changeset
   148
*}
2649
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   149
2651
4aa72a88b2c1 equivariance of THE_default under the uniqueness assumption
Christian Urban <urbanc@in.tum.de>
parents: 2650
diff changeset
   150
thm accp_induct_rule
2649
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   151
2651
4aa72a88b2c1 equivariance of THE_default under the uniqueness assumption
Christian Urban <urbanc@in.tum.de>
parents: 2650
diff changeset
   152
ML {*
2649
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   153
(* Theory dependencies. *)
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   154
val acc_induct_rule = @{thm accp_induct_rule}
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   155
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   156
val ex1_implies_ex = @{thm FunDef.fundef_ex1_existence}
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   157
val ex1_implies_un = @{thm FunDef.fundef_ex1_uniqueness}
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   158
val ex1_implies_iff = @{thm FunDef.fundef_ex1_iff}
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   159
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   160
val acc_downward = @{thm accp_downward}
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   161
val accI = @{thm accp.accI}
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   162
val case_split = @{thm HOL.case_split}
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   163
val fundef_default_value = @{thm FunDef.fundef_default_value}
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   164
val not_acc_down = @{thm not_accp_down}
2660
3342a2d13d95 nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
parents: 2657
diff changeset
   165
*}
2649
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   166
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   167
2660
3342a2d13d95 nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
parents: 2657
diff changeset
   168
ML {*
2649
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   169
fun find_calls tree =
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   170
  let
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   171
    fun add_Ri (fixes,assumes) (_ $ arg) _ (_, xs) =
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   172
      ([], (fixes, assumes, arg) :: xs)
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   173
      | add_Ri _ _ _ _ = raise Match
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   174
  in
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   175
    rev (Function_Ctx_Tree.traverse_tree add_Ri tree [])
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   176
  end
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   177
*}
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   178
2650
e5fa8de0e4bd derived equivariance for the function graph and function relation
Christian Urban <urbanc@in.tum.de>
parents: 2649
diff changeset
   179
ML {*
2653
d0f774d06e6e added eqvt_at premises in function definition - however not proved at the moment
Christian Urban <urbanc@in.tum.de>
parents: 2652
diff changeset
   180
fun mk_eqvt_at (f_trm, arg_trm) =
2650
e5fa8de0e4bd derived equivariance for the function graph and function relation
Christian Urban <urbanc@in.tum.de>
parents: 2649
diff changeset
   181
  let
2653
d0f774d06e6e added eqvt_at premises in function definition - however not proved at the moment
Christian Urban <urbanc@in.tum.de>
parents: 2652
diff changeset
   182
    val f_ty = fastype_of f_trm
d0f774d06e6e added eqvt_at premises in function definition - however not proved at the moment
Christian Urban <urbanc@in.tum.de>
parents: 2652
diff changeset
   183
    val arg_ty = domain_type f_ty
2650
e5fa8de0e4bd derived equivariance for the function graph and function relation
Christian Urban <urbanc@in.tum.de>
parents: 2649
diff changeset
   184
  in
2653
d0f774d06e6e added eqvt_at premises in function definition - however not proved at the moment
Christian Urban <urbanc@in.tum.de>
parents: 2652
diff changeset
   185
    Const (@{const_name eqvt_at}, [f_ty, arg_ty] ---> @{typ bool}) $ f_trm $ arg_trm
2650
e5fa8de0e4bd derived equivariance for the function graph and function relation
Christian Urban <urbanc@in.tum.de>
parents: 2649
diff changeset
   186
    |> HOLogic.mk_Trueprop
e5fa8de0e4bd derived equivariance for the function graph and function relation
Christian Urban <urbanc@in.tum.de>
parents: 2649
diff changeset
   187
  end
e5fa8de0e4bd derived equivariance for the function graph and function relation
Christian Urban <urbanc@in.tum.de>
parents: 2649
diff changeset
   188
*}
e5fa8de0e4bd derived equivariance for the function graph and function relation
Christian Urban <urbanc@in.tum.de>
parents: 2649
diff changeset
   189
2653
d0f774d06e6e added eqvt_at premises in function definition - however not proved at the moment
Christian Urban <urbanc@in.tum.de>
parents: 2652
diff changeset
   190
ML {*
d0f774d06e6e added eqvt_at premises in function definition - however not proved at the moment
Christian Urban <urbanc@in.tum.de>
parents: 2652
diff changeset
   191
fun find_calls2 f t = 
d0f774d06e6e added eqvt_at premises in function definition - however not proved at the moment
Christian Urban <urbanc@in.tum.de>
parents: 2652
diff changeset
   192
  let
d0f774d06e6e added eqvt_at premises in function definition - however not proved at the moment
Christian Urban <urbanc@in.tum.de>
parents: 2652
diff changeset
   193
    fun aux (g $ arg) = aux g #> aux arg #> (if f = g then cons ((g, arg)) else I)
d0f774d06e6e added eqvt_at premises in function definition - however not proved at the moment
Christian Urban <urbanc@in.tum.de>
parents: 2652
diff changeset
   194
      | aux (Abs (_, _, t)) = aux t 
d0f774d06e6e added eqvt_at premises in function definition - however not proved at the moment
Christian Urban <urbanc@in.tum.de>
parents: 2652
diff changeset
   195
      | aux _ = I
d0f774d06e6e added eqvt_at premises in function definition - however not proved at the moment
Christian Urban <urbanc@in.tum.de>
parents: 2652
diff changeset
   196
  in
d0f774d06e6e added eqvt_at premises in function definition - however not proved at the moment
Christian Urban <urbanc@in.tum.de>
parents: 2652
diff changeset
   197
    aux t []
d0f774d06e6e added eqvt_at premises in function definition - however not proved at the moment
Christian Urban <urbanc@in.tum.de>
parents: 2652
diff changeset
   198
  end 
d0f774d06e6e added eqvt_at premises in function definition - however not proved at the moment
Christian Urban <urbanc@in.tum.de>
parents: 2652
diff changeset
   199
*}
d0f774d06e6e added eqvt_at premises in function definition - however not proved at the moment
Christian Urban <urbanc@in.tum.de>
parents: 2652
diff changeset
   200
2649
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   201
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   202
ML {*
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   203
(** building proof obligations *)
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   204
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   205
fun mk_compat_proof_obligations domT ranT fvar f glrs =
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   206
  let
2653
d0f774d06e6e added eqvt_at premises in function definition - however not proved at the moment
Christian Urban <urbanc@in.tum.de>
parents: 2652
diff changeset
   207
    fun mk_impl ((qs, gs, lhs, rhs), (qs', gs', lhs', rhs')) =
2649
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   208
      let
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   209
        val shift = incr_boundvars (length qs')
2653
d0f774d06e6e added eqvt_at premises in function definition - however not proved at the moment
Christian Urban <urbanc@in.tum.de>
parents: 2652
diff changeset
   210
d0f774d06e6e added eqvt_at premises in function definition - however not proved at the moment
Christian Urban <urbanc@in.tum.de>
parents: 2652
diff changeset
   211
        val RCs_rhs  = find_calls2 fvar rhs
d0f774d06e6e added eqvt_at premises in function definition - however not proved at the moment
Christian Urban <urbanc@in.tum.de>
parents: 2652
diff changeset
   212
        val RCs_rhs' = find_calls2 fvar rhs'
2649
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   213
      in
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   214
        Logic.mk_implies
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   215
          (HOLogic.mk_Trueprop (HOLogic.eq_const domT $ shift lhs $ lhs'),
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   216
            HOLogic.mk_Trueprop (HOLogic.eq_const ranT $ shift rhs $ rhs'))
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   217
        |> fold_rev (curry Logic.mk_implies) (map shift gs @ gs')
2660
3342a2d13d95 nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
parents: 2657
diff changeset
   218
        |> fold_rev (curry Logic.mk_implies) (map (shift o mk_eqvt_at) RCs_rhs) (* HERE *) 
3342a2d13d95 nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
parents: 2657
diff changeset
   219
        (*|> fold_rev (curry Logic.mk_implies) (map mk_eqvt_at RCs_rhs')*) (* HERE *) 
2649
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   220
        |> fold_rev (fn (n,T) => fn b => Term.all T $ Abs(n,T,b)) (qs @ qs')
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   221
        |> curry abstract_over fvar
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   222
        |> curry subst_bound f
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   223
      end
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   224
  in
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   225
    map mk_impl (unordered_pairs glrs)
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   226
  end
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   227
*}
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   228
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   229
ML {*
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   230
fun mk_completeness (Globals {x, Pbool, ...}) clauses qglrs =
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   231
  let
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   232
    fun mk_case (ClauseContext {qs, gs, lhs, ...}, (oqs, _, _, _)) =
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   233
      HOLogic.mk_Trueprop Pbool
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   234
      |> curry Logic.mk_implies (HOLogic.mk_Trueprop (mk_eq (x, lhs)))
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   235
      |> fold_rev (curry Logic.mk_implies) gs
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   236
      |> fold_rev mk_forall_rename (map fst oqs ~~ qs)
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   237
  in
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   238
    HOLogic.mk_Trueprop Pbool
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   239
    |> fold_rev (curry Logic.mk_implies o mk_case) (clauses ~~ qglrs)
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   240
    |> mk_forall_rename ("x", x)
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   241
    |> mk_forall_rename ("P", Pbool)
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   242
  end
2660
3342a2d13d95 nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
parents: 2657
diff changeset
   243
*}
2649
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   244
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   245
(** making a context with it's own local bindings **)
2660
3342a2d13d95 nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
parents: 2657
diff changeset
   246
ML {*
2649
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   247
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   248
fun mk_clause_context x ctxt (pre_qs,pre_gs,pre_lhs,pre_rhs) =
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   249
  let
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   250
    val (qs, ctxt') = Variable.variant_fixes (map fst pre_qs) ctxt
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   251
      |>> map2 (fn (_, T) => fn n => Free (n, T)) pre_qs
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   252
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   253
    val thy = ProofContext.theory_of ctxt'
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   254
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   255
    fun inst t = subst_bounds (rev qs, t)
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   256
    val gs = map inst pre_gs
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   257
    val lhs = inst pre_lhs
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   258
    val rhs = inst pre_rhs
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   259
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   260
    val cqs = map (cterm_of thy) qs
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   261
    val ags = map (Thm.assume o cterm_of thy) gs
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   262
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   263
    val case_hyp = Thm.assume (cterm_of thy (HOLogic.mk_Trueprop (mk_eq (x, lhs))))
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   264
  in
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   265
    ClauseContext { ctxt = ctxt', qs = qs, gs = gs, lhs = lhs, rhs = rhs,
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   266
      cqs = cqs, ags = ags, case_hyp = case_hyp }
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   267
  end
2660
3342a2d13d95 nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
parents: 2657
diff changeset
   268
*}
2649
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   269
2660
3342a2d13d95 nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
parents: 2657
diff changeset
   270
ML {*
2649
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   271
(* lowlevel term function. FIXME: remove *)
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   272
fun abstract_over_list vs body =
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   273
  let
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   274
    fun abs lev v tm =
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   275
      if v aconv tm then Bound lev
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   276
      else
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   277
        (case tm of
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   278
          Abs (a, T, t) => Abs (a, T, abs (lev + 1) v t)
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   279
        | t $ u => abs lev v t $ abs lev v u
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   280
        | t => t)
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   281
  in
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   282
    fold_index (fn (i, v) => fn t => abs i v t) vs body
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   283
  end
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   284
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   285
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   286
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   287
fun mk_clause_info globals G f no cdata qglr tree RCs GIntro_thm RIntro_thms =
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   288
  let
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   289
    val Globals {h, ...} = globals
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   290
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   291
    val ClauseContext { ctxt, qs, cqs, ags, ... } = cdata
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   292
    val cert = Thm.cterm_of (ProofContext.theory_of ctxt)
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   293
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   294
    (* Instantiate the GIntro thm with "f" and import into the clause context. *)
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   295
    val lGI = GIntro_thm
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   296
      |> Thm.forall_elim (cert f)
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   297
      |> fold Thm.forall_elim cqs
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   298
      |> fold Thm.elim_implies ags
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   299
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   300
    fun mk_call_info (rcfix, rcassm, rcarg) RI =
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   301
      let
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   302
        val llRI = RI
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   303
          |> fold Thm.forall_elim cqs
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   304
          |> fold (Thm.forall_elim o cert o Free) rcfix
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   305
          |> fold Thm.elim_implies ags
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   306
          |> fold Thm.elim_implies rcassm
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   307
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   308
        val h_assum =
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   309
          HOLogic.mk_Trueprop (G $ rcarg $ (h $ rcarg))
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   310
          |> fold_rev (curry Logic.mk_implies o prop_of) rcassm
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   311
          |> fold_rev (Logic.all o Free) rcfix
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   312
          |> Pattern.rewrite_term (ProofContext.theory_of ctxt) [(f, h)] []
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   313
          |> abstract_over_list (rev qs)
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   314
      in
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   315
        RCInfo {RIvs=rcfix, rcarg=rcarg, CCas=rcassm, llRI=llRI, h_assum=h_assum}
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   316
      end
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   317
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   318
    val RC_infos = map2 mk_call_info RCs RIntro_thms
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   319
  in
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   320
    ClauseInfo {no=no, cdata=cdata, qglr=qglr, lGI=lGI, RCs=RC_infos,
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   321
      tree=tree}
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   322
  end
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   323
*}
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   324
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   325
ML {*
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   326
fun store_compat_thms 0 thms = []
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   327
  | store_compat_thms n thms =
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   328
  let
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   329
    val (thms1, thms2) = chop n thms
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   330
  in
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   331
    (thms1 :: store_compat_thms (n - 1) thms2)
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   332
  end
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   333
*}
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   334
2660
3342a2d13d95 nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
parents: 2657
diff changeset
   335
ML {*
3342a2d13d95 nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
parents: 2657
diff changeset
   336
fun pp_thm thm =
3342a2d13d95 nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
parents: 2657
diff changeset
   337
  let
3342a2d13d95 nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
parents: 2657
diff changeset
   338
    val hyps = Thm.hyps_of thm
3342a2d13d95 nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
parents: 2657
diff changeset
   339
    val tpairs = Thm.terms_of_tpairs (Thm.tpairs_of thm)
3342a2d13d95 nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
parents: 2657
diff changeset
   340
  in
3342a2d13d95 nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
parents: 2657
diff changeset
   341
      (@{make_string} thm) ^ "\n    hyps: " ^ (commas (map (Syntax.string_of_term @{context}) hyps))
3342a2d13d95 nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
parents: 2657
diff changeset
   342
      ^ "    tpairs: " ^   (commas (map (Syntax.string_of_term @{context}) tpairs))
3342a2d13d95 nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
parents: 2657
diff changeset
   343
  end
3342a2d13d95 nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
parents: 2657
diff changeset
   344
*}
3342a2d13d95 nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
parents: 2657
diff changeset
   345
2653
d0f774d06e6e added eqvt_at premises in function definition - however not proved at the moment
Christian Urban <urbanc@in.tum.de>
parents: 2652
diff changeset
   346
d0f774d06e6e added eqvt_at premises in function definition - however not proved at the moment
Christian Urban <urbanc@in.tum.de>
parents: 2652
diff changeset
   347
ML {*
2660
3342a2d13d95 nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
parents: 2657
diff changeset
   348
fun eqvt_at_elim thy (eqvts:thm list) thm =
2653
d0f774d06e6e added eqvt_at premises in function definition - however not proved at the moment
Christian Urban <urbanc@in.tum.de>
parents: 2652
diff changeset
   349
  case (prop_of thm) of
d0f774d06e6e added eqvt_at premises in function definition - however not proved at the moment
Christian Urban <urbanc@in.tum.de>
parents: 2652
diff changeset
   350
    Const ("==>", _) $ (t as (@{term Trueprop} $ (Const (@{const_name eqvt_at}, _) $ _ $ _))) $ _ => 
d0f774d06e6e added eqvt_at premises in function definition - however not proved at the moment
Christian Urban <urbanc@in.tum.de>
parents: 2652
diff changeset
   351
      let
2660
3342a2d13d95 nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
parents: 2657
diff changeset
   352
        val el_thm = Skip_Proof.make_thm thy t
3342a2d13d95 nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
parents: 2657
diff changeset
   353
        val _ = tracing ("NEED TO PROVE  " ^ @{make_string} el_thm)
3342a2d13d95 nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
parents: 2657
diff changeset
   354
        val _ = tracing ("HAVE\n" ^ cat_lines (map @{make_string} eqvts))
2653
d0f774d06e6e added eqvt_at premises in function definition - however not proved at the moment
Christian Urban <urbanc@in.tum.de>
parents: 2652
diff changeset
   355
      in
d0f774d06e6e added eqvt_at premises in function definition - however not proved at the moment
Christian Urban <urbanc@in.tum.de>
parents: 2652
diff changeset
   356
        Thm.elim_implies el_thm thm 
2660
3342a2d13d95 nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
parents: 2657
diff changeset
   357
        |> eqvt_at_elim thy eqvts 
2653
d0f774d06e6e added eqvt_at premises in function definition - however not proved at the moment
Christian Urban <urbanc@in.tum.de>
parents: 2652
diff changeset
   358
      end
d0f774d06e6e added eqvt_at premises in function definition - however not proved at the moment
Christian Urban <urbanc@in.tum.de>
parents: 2652
diff changeset
   359
  | _ => thm
d0f774d06e6e added eqvt_at premises in function definition - however not proved at the moment
Christian Urban <urbanc@in.tum.de>
parents: 2652
diff changeset
   360
*}
d0f774d06e6e added eqvt_at premises in function definition - however not proved at the moment
Christian Urban <urbanc@in.tum.de>
parents: 2652
diff changeset
   361
2649
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   362
ML {*
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   363
(* expects i <= j *)
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   364
fun lookup_compat_thm i j cts =
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   365
  nth (nth cts (i - 1)) (j - i)
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   366
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   367
(* Returns "Gsi, Gsj, lhs_i = lhs_j |-- rhs_j_f = rhs_i_f" *)
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   368
(* if j < i, then turn around *)
2660
3342a2d13d95 nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
parents: 2657
diff changeset
   369
fun get_compat_thm thy cts eqvtsi eqvtsj i j ctxi ctxj =
2649
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   370
  let
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   371
    val ClauseContext {cqs=cqsi,ags=agsi,lhs=lhsi,...} = ctxi
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   372
    val ClauseContext {cqs=cqsj,ags=agsj,lhs=lhsj,...} = ctxj
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   373
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   374
    val lhsi_eq_lhsj = cterm_of thy (HOLogic.mk_Trueprop (mk_eq (lhsi, lhsj)))
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   375
  in if j < i then
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   376
    let
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   377
      val compat = lookup_compat_thm j i cts
2660
3342a2d13d95 nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
parents: 2657
diff changeset
   378
      val _ = tracing ("XXXX compat clause if:\n" ^ @{make_string} compat)
2649
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   379
    in
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   380
      compat         (* "!!qj qi. Gsj => Gsi => lhsj = lhsi ==> rhsj = rhsi" *)
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   381
      |> fold Thm.forall_elim (cqsj @ cqsi) (* "Gsj => Gsi => lhsj = lhsi ==> rhsj = rhsi" *)
2660
3342a2d13d95 nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
parents: 2657
diff changeset
   382
      |> fold Thm.elim_implies (rev eqvtsj) (* HERE *)
3342a2d13d95 nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
parents: 2657
diff changeset
   383
      (*|> eqvt_at_elim thy eqvtsj *)
2649
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   384
      |> fold Thm.elim_implies agsj
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   385
      |> fold Thm.elim_implies agsi
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   386
      |> Thm.elim_implies ((Thm.assume lhsi_eq_lhsj) RS sym) (* "Gsj, Gsi, lhsi = lhsj |-- rhsj = rhsi" *)
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   387
    end
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   388
    else
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   389
    let
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   390
      val compat = lookup_compat_thm i j cts
2660
3342a2d13d95 nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
parents: 2657
diff changeset
   391
      (* val _ = tracing ("XXXX compat clause else:\n" ^ @{make_string} compat) *)
2649
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   392
    in
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   393
      compat        (* "!!qi qj. Gsi => Gsj => lhsi = lhsj ==> rhsi = rhsj" *)
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   394
      |> fold Thm.forall_elim (cqsi @ cqsj) (* "Gsi => Gsj => lhsi = lhsj ==> rhsi = rhsj" *)
2661
16896fd8eff5 subst also works now
Christian Urban <urbanc@in.tum.de>
parents: 2660
diff changeset
   395
      |> fold Thm.elim_implies (rev eqvtsi)  (* HERE *)
16896fd8eff5 subst also works now
Christian Urban <urbanc@in.tum.de>
parents: 2660
diff changeset
   396
      (* |> eqvt_at_elim thy eqvtsi *)
2649
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   397
      |> fold Thm.elim_implies agsi
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   398
      |> fold Thm.elim_implies agsj
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   399
      |> Thm.elim_implies (Thm.assume lhsi_eq_lhsj)
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   400
      |> (fn thm => thm RS sym) (* "Gsi, Gsj, lhsi = lhsj |-- rhsj = rhsi" *)
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   401
    end
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   402
  end
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   403
*}
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   404
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   405
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   406
ML {*
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   407
(* Generates the replacement lemma in fully quantified form. *)
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   408
fun mk_replacement_lemma thy h ih_elim clause =
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   409
  let
2660
3342a2d13d95 nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
parents: 2657
diff changeset
   410
    val ClauseInfo {cdata=ClauseContext {qs, cqs, ags, case_hyp, ...},
2649
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   411
      RCs, tree, ...} = clause
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   412
    local open Conv in
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   413
      val ih_conv = arg1_conv o arg_conv o arg_conv
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   414
    end
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   415
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   416
    val ih_elim_case =
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   417
      Conv.fconv_rule (ih_conv (K (case_hyp RS eq_reflection))) ih_elim
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   418
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   419
    val Ris = map (fn RCInfo {llRI, ...} => llRI) RCs
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   420
    val h_assums = map (fn RCInfo {h_assum, ...} =>
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   421
      Thm.assume (cterm_of thy (subst_bounds (rev qs, h_assum)))) RCs
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   422
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   423
    val (eql, _) =
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   424
      Function_Ctx_Tree.rewrite_by_tree thy h ih_elim_case (Ris ~~ h_assums) tree
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   425
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   426
    val replace_lemma = (eql RS meta_eq_to_obj_eq)
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   427
      |> Thm.implies_intr (cprop_of case_hyp)
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   428
      |> fold_rev (Thm.implies_intr o cprop_of) h_assums
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   429
      |> fold_rev (Thm.implies_intr o cprop_of) ags
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   430
      |> fold_rev Thm.forall_intr cqs
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   431
      |> Thm.close_derivation
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   432
  in
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   433
    replace_lemma
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   434
  end
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   435
*}
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   436
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   437
ML {*
2660
3342a2d13d95 nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
parents: 2657
diff changeset
   438
fun mk_eqvt_lemma thy ih_eqvt clause =
3342a2d13d95 nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
parents: 2657
diff changeset
   439
  let
2661
16896fd8eff5 subst also works now
Christian Urban <urbanc@in.tum.de>
parents: 2660
diff changeset
   440
    val ClauseInfo {cdata=ClauseContext {cqs, ags, case_hyp, ...}, RCs, ...} = clause
2660
3342a2d13d95 nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
parents: 2657
diff changeset
   441
3342a2d13d95 nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
parents: 2657
diff changeset
   442
    local open Conv in
3342a2d13d95 nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
parents: 2657
diff changeset
   443
      val ih_conv = arg1_conv o arg_conv o arg_conv
3342a2d13d95 nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
parents: 2657
diff changeset
   444
    end
3342a2d13d95 nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
parents: 2657
diff changeset
   445
3342a2d13d95 nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
parents: 2657
diff changeset
   446
    val ih_eqvt_case =
3342a2d13d95 nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
parents: 2657
diff changeset
   447
      Conv.fconv_rule (ih_conv (K (case_hyp RS eq_reflection))) ih_eqvt
3342a2d13d95 nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
parents: 2657
diff changeset
   448
3342a2d13d95 nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
parents: 2657
diff changeset
   449
    fun prep_eqvt (RCInfo {llRI, RIvs, CCas, ...}) = 
3342a2d13d95 nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
parents: 2657
diff changeset
   450
        (llRI RS ih_eqvt_case)
3342a2d13d95 nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
parents: 2657
diff changeset
   451
        |> fold_rev (Thm.implies_intr o cprop_of) CCas
3342a2d13d95 nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
parents: 2657
diff changeset
   452
        |> fold_rev (Thm.forall_intr o cterm_of thy o Free) RIvs
3342a2d13d95 nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
parents: 2657
diff changeset
   453
  in
3342a2d13d95 nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
parents: 2657
diff changeset
   454
    map prep_eqvt RCs
2661
16896fd8eff5 subst also works now
Christian Urban <urbanc@in.tum.de>
parents: 2660
diff changeset
   455
    |> map (fold_rev (Thm.implies_intr o cprop_of) ags)
2660
3342a2d13d95 nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
parents: 2657
diff changeset
   456
    |> map (Thm.implies_intr (cprop_of case_hyp))
3342a2d13d95 nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
parents: 2657
diff changeset
   457
    |> map (fold_rev Thm.forall_intr cqs)
3342a2d13d95 nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
parents: 2657
diff changeset
   458
    |> map (Thm.close_derivation) 
3342a2d13d95 nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
parents: 2657
diff changeset
   459
  end
3342a2d13d95 nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
parents: 2657
diff changeset
   460
*}
3342a2d13d95 nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
parents: 2657
diff changeset
   461
3342a2d13d95 nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
parents: 2657
diff changeset
   462
ML {*
3342a2d13d95 nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
parents: 2657
diff changeset
   463
fun mk_uniqueness_clause thy globals compat_store eqvts clausei clausej RLj =
2649
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   464
  let
2661
16896fd8eff5 subst also works now
Christian Urban <urbanc@in.tum.de>
parents: 2660
diff changeset
   465
    val _ = tracing "START"
16896fd8eff5 subst also works now
Christian Urban <urbanc@in.tum.de>
parents: 2660
diff changeset
   466
2649
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   467
    val Globals {h, y, x, fvar, ...} = globals
2661
16896fd8eff5 subst also works now
Christian Urban <urbanc@in.tum.de>
parents: 2660
diff changeset
   468
    val ClauseInfo {no=i, cdata=cctxi as ClauseContext {ctxt=ctxti, lhs=lhsi, case_hyp, cqs = cqsi, ags = agsi, ...}, ...} = clausei
2649
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   469
    val ClauseInfo {no=j, qglr=cdescj, RCs=RCsj, ...} = clausej
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   470
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   471
    val cctxj as ClauseContext {ags = agsj', lhs = lhsj', rhs = rhsj', qs = qsj', cqs = cqsj', ...} =
2660
3342a2d13d95 nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
parents: 2657
diff changeset
   472
      mk_clause_context x ctxti cdescj 
2649
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   473
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   474
    val rhsj'h = Pattern.rewrite_term thy [(fvar,h)] [] rhsj'
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   475
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   476
    val Ghsj' = map 
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   477
      (fn RCInfo {h_assum, ...} => Thm.assume (cterm_of thy (subst_bounds (rev qsj', h_assum)))) RCsj
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   478
2660
3342a2d13d95 nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
parents: 2657
diff changeset
   479
    val y_eq_rhsj'h = Thm.assume (cterm_of thy (HOLogic.mk_Trueprop (mk_eq (y, rhsj'h))))
3342a2d13d95 nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
parents: 2657
diff changeset
   480
    val lhsi_eq_lhsj' = Thm.assume (cterm_of thy (HOLogic.mk_Trueprop (mk_eq (lhsi, lhsj'))))
3342a2d13d95 nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
parents: 2657
diff changeset
   481
       (* lhs_i = lhs_j' |-- lhs_i = lhs_j' *) 
3342a2d13d95 nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
parents: 2657
diff changeset
   482
2661
16896fd8eff5 subst also works now
Christian Urban <urbanc@in.tum.de>
parents: 2660
diff changeset
   483
    val _ = tracing "POINT B"
16896fd8eff5 subst also works now
Christian Urban <urbanc@in.tum.de>
parents: 2660
diff changeset
   484
2660
3342a2d13d95 nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
parents: 2657
diff changeset
   485
    val case_hypj' = trans OF [case_hyp, lhsi_eq_lhsj']
3342a2d13d95 nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
parents: 2657
diff changeset
   486
2649
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   487
    val RLj_import = RLj
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   488
      |> fold Thm.forall_elim cqsj'
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   489
      |> fold Thm.elim_implies agsj'
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   490
      |> fold Thm.elim_implies Ghsj'
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   491
2661
16896fd8eff5 subst also works now
Christian Urban <urbanc@in.tum.de>
parents: 2660
diff changeset
   492
    val _ = tracing "POINT C"
16896fd8eff5 subst also works now
Christian Urban <urbanc@in.tum.de>
parents: 2660
diff changeset
   493
 
2660
3342a2d13d95 nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
parents: 2657
diff changeset
   494
    val eqvtsi = nth eqvts (i - 1)
3342a2d13d95 nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
parents: 2657
diff changeset
   495
      |> map (fold Thm.forall_elim cqsi)
3342a2d13d95 nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
parents: 2657
diff changeset
   496
      |> map (fold Thm.elim_implies [case_hyp])
2661
16896fd8eff5 subst also works now
Christian Urban <urbanc@in.tum.de>
parents: 2660
diff changeset
   497
      |> map (fold Thm.elim_implies agsi)
16896fd8eff5 subst also works now
Christian Urban <urbanc@in.tum.de>
parents: 2660
diff changeset
   498
16896fd8eff5 subst also works now
Christian Urban <urbanc@in.tum.de>
parents: 2660
diff changeset
   499
    val _ = tracing "POINT D"
2660
3342a2d13d95 nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
parents: 2657
diff changeset
   500
3342a2d13d95 nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
parents: 2657
diff changeset
   501
    val eqvtsj = nth eqvts (j - 1)
2661
16896fd8eff5 subst also works now
Christian Urban <urbanc@in.tum.de>
parents: 2660
diff changeset
   502
      |> tap (fn thms => tracing ("O1:\n" ^ cat_lines (map @{make_string} thms)))
2660
3342a2d13d95 nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
parents: 2657
diff changeset
   503
      |> map (fold Thm.forall_elim cqsj')
2661
16896fd8eff5 subst also works now
Christian Urban <urbanc@in.tum.de>
parents: 2660
diff changeset
   504
      |> tap (fn thms => tracing ("O2:\n" ^ cat_lines (map @{make_string} thms)))
2660
3342a2d13d95 nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
parents: 2657
diff changeset
   505
      |> map (fold Thm.elim_implies [case_hypj'])
2661
16896fd8eff5 subst also works now
Christian Urban <urbanc@in.tum.de>
parents: 2660
diff changeset
   506
      |> tap (fn thms => tracing ("O3:\n" ^ cat_lines (map @{make_string} thms)))
16896fd8eff5 subst also works now
Christian Urban <urbanc@in.tum.de>
parents: 2660
diff changeset
   507
      |> map (fold Thm.elim_implies agsj')
2660
3342a2d13d95 nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
parents: 2657
diff changeset
   508
3342a2d13d95 nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
parents: 2657
diff changeset
   509
    val _ = tracing ("eqvtsi:\n" ^ cat_lines (map @{make_string} eqvtsi))
3342a2d13d95 nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
parents: 2657
diff changeset
   510
    val _ = tracing ("eqvtsj:\n" ^ cat_lines (map @{make_string} eqvtsj))
3342a2d13d95 nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
parents: 2657
diff changeset
   511
    val _ = tracing ("case_hypi:\n" ^ @{make_string} case_hyp)
3342a2d13d95 nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
parents: 2657
diff changeset
   512
    val _ = tracing ("case_hypj:\n" ^ @{make_string} case_hypj')
3342a2d13d95 nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
parents: 2657
diff changeset
   513
3342a2d13d95 nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
parents: 2657
diff changeset
   514
    val compat = get_compat_thm thy compat_store eqvtsi eqvtsj i j cctxi cctxj
3342a2d13d95 nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
parents: 2657
diff changeset
   515
    
2661
16896fd8eff5 subst also works now
Christian Urban <urbanc@in.tum.de>
parents: 2660
diff changeset
   516
    val _ = tracing ("compats:\n" ^ pp_thm compat) 
2660
3342a2d13d95 nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
parents: 2657
diff changeset
   517
    
3342a2d13d95 nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
parents: 2657
diff changeset
   518
3342a2d13d95 nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
parents: 2657
diff changeset
   519
    fun pppp str = tap (fn thm => tracing (str ^ ": " ^ pp_thm thm))
3342a2d13d95 nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
parents: 2657
diff changeset
   520
    fun ppp str = I
2649
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   521
  in
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   522
    (trans OF [case_hyp, lhsi_eq_lhsj']) (* lhs_i = lhs_j' |-- x = lhs_j' *)
2661
16896fd8eff5 subst also works now
Christian Urban <urbanc@in.tum.de>
parents: 2660
diff changeset
   523
    |> pppp "A"
2649
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   524
    |> Thm.implies_elim RLj_import
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   525
      (* Rj1' ... Rjk', lhs_i = lhs_j' |-- rhs_j'_h = rhs_j'_f *)
2661
16896fd8eff5 subst also works now
Christian Urban <urbanc@in.tum.de>
parents: 2660
diff changeset
   526
    |> pppp "B"
2649
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   527
    |> (fn it => trans OF [it, compat])
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   528
      (* lhs_i = lhs_j', Gj', Rj1' ... Rjk' |-- rhs_j'_h = rhs_i_f *)
2661
16896fd8eff5 subst also works now
Christian Urban <urbanc@in.tum.de>
parents: 2660
diff changeset
   529
    |> pppp "C"
2649
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   530
    |> (fn it => trans OF [y_eq_rhsj'h, it])
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   531
      (* lhs_i = lhs_j', Gj', Rj1' ... Rjk', y = rhs_j_h' |-- y = rhs_i_f *)
2661
16896fd8eff5 subst also works now
Christian Urban <urbanc@in.tum.de>
parents: 2660
diff changeset
   532
    |> pppp "D"
2649
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   533
    |> fold_rev (Thm.implies_intr o cprop_of) Ghsj'
2661
16896fd8eff5 subst also works now
Christian Urban <urbanc@in.tum.de>
parents: 2660
diff changeset
   534
    |> pppp "E"
2649
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   535
    |> fold_rev (Thm.implies_intr o cprop_of) agsj'
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   536
      (* lhs_i = lhs_j' , y = rhs_j_h' |-- Gj', Rj1'...Rjk' ==> y = rhs_i_f *)
2661
16896fd8eff5 subst also works now
Christian Urban <urbanc@in.tum.de>
parents: 2660
diff changeset
   537
    |> pppp "F"
2649
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   538
    |> Thm.implies_intr (cprop_of y_eq_rhsj'h)
2661
16896fd8eff5 subst also works now
Christian Urban <urbanc@in.tum.de>
parents: 2660
diff changeset
   539
    |> pppp "G"
2649
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   540
    |> Thm.implies_intr (cprop_of lhsi_eq_lhsj')
2661
16896fd8eff5 subst also works now
Christian Urban <urbanc@in.tum.de>
parents: 2660
diff changeset
   541
    |> pppp "H"
2649
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   542
    |> fold_rev Thm.forall_intr (cterm_of thy h :: cqsj')
2661
16896fd8eff5 subst also works now
Christian Urban <urbanc@in.tum.de>
parents: 2660
diff changeset
   543
    |> pppp "I"
2649
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   544
  end
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   545
*}
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   546
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   547
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   548
ML {*
2660
3342a2d13d95 nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
parents: 2657
diff changeset
   549
fun mk_uniqueness_case thy globals G f ihyp ih_intro G_cases compat_store clauses rep_lemmas eqvt_lemmas clausei =
2649
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   550
  let
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   551
    val Globals {x, y, ranT, fvar, ...} = globals
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   552
    val ClauseInfo {cdata = ClauseContext {lhs, rhs, cqs, ags, case_hyp, ...}, lGI, RCs, ...} = clausei
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   553
    val rhsC = Pattern.rewrite_term thy [(fvar, f)] [] rhs
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   554
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   555
    val ih_intro_case = full_simplify (HOL_basic_ss addsimps [case_hyp]) ih_intro
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   556
2660
3342a2d13d95 nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
parents: 2657
diff changeset
   557
    fun prep_RC (RCInfo {llRI, RIvs, CCas, ...}) = 
3342a2d13d95 nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
parents: 2657
diff changeset
   558
        (llRI RS ih_intro_case)
3342a2d13d95 nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
parents: 2657
diff changeset
   559
        |> fold_rev (Thm.implies_intr o cprop_of) CCas
3342a2d13d95 nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
parents: 2657
diff changeset
   560
        |> fold_rev (Thm.forall_intr o cterm_of thy o Free) RIvs
2649
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   561
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   562
    val existence = fold (curry op COMP o prep_RC) RCs lGI
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   563
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   564
    val P = cterm_of thy (mk_eq (y, rhsC))
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   565
    val G_lhs_y = Thm.assume (cterm_of thy (HOLogic.mk_Trueprop (G $ lhs $ y)))
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   566
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   567
    val unique_clauses =
2660
3342a2d13d95 nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
parents: 2657
diff changeset
   568
      map2 (mk_uniqueness_clause thy globals compat_store eqvt_lemmas clausei) clauses rep_lemmas
2651
4aa72a88b2c1 equivariance of THE_default under the uniqueness assumption
Christian Urban <urbanc@in.tum.de>
parents: 2650
diff changeset
   569
2661
16896fd8eff5 subst also works now
Christian Urban <urbanc@in.tum.de>
parents: 2660
diff changeset
   570
    val _ = tracing ("DONE unique clauses") 
16896fd8eff5 subst also works now
Christian Urban <urbanc@in.tum.de>
parents: 2660
diff changeset
   571
2649
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   572
    fun elim_implies_eta A AB =
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   573
      Thm.compose_no_flatten true (A, 0) 1 AB |> Seq.list_of |> the_single
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   574
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   575
    val uniqueness = G_cases
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   576
      |> Thm.forall_elim (cterm_of thy lhs)
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   577
      |> Thm.forall_elim (cterm_of thy y)
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   578
      |> Thm.forall_elim P
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   579
      |> Thm.elim_implies G_lhs_y
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   580
      |> fold elim_implies_eta unique_clauses
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   581
      |> Thm.implies_intr (cprop_of G_lhs_y)
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   582
      |> Thm.forall_intr (cterm_of thy y)
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   583
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   584
    val P2 = cterm_of thy (lambda y (G $ lhs $ y)) (* P2 y := (lhs, y): G *)
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   585
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   586
    val exactly_one =
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   587
      ex1I |> instantiate' [SOME (ctyp_of thy ranT)] [SOME P2, SOME (cterm_of thy rhsC)]
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   588
      |> curry (op COMP) existence
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   589
      |> curry (op COMP) uniqueness
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   590
      |> simplify (HOL_basic_ss addsimps [case_hyp RS sym])
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   591
      |> Thm.implies_intr (cprop_of case_hyp)
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   592
      |> fold_rev (Thm.implies_intr o cprop_of) ags
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   593
      |> fold_rev Thm.forall_intr cqs
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   594
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   595
    val function_value =
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   596
      existence
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   597
      |> Thm.implies_intr ihyp
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   598
      |> Thm.implies_intr (cprop_of case_hyp)
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   599
      |> Thm.forall_intr (cterm_of thy x)
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   600
      |> Thm.forall_elim (cterm_of thy lhs)
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   601
      |> curry (op RS) refl
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   602
  in
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   603
    (exactly_one, function_value)
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   604
  end
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   605
*}
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   606
2655
1c3ad1256f16 instantiated fundef_ex1_eqvt_at theorem with the indction hypothesis
Christian Urban <urbanc@in.tum.de>
parents: 2654
diff changeset
   607
1c3ad1256f16 instantiated fundef_ex1_eqvt_at theorem with the indction hypothesis
Christian Urban <urbanc@in.tum.de>
parents: 2654
diff changeset
   608
ML {* (ex1_implies_ex, ex1_implies_un) *}
1c3ad1256f16 instantiated fundef_ex1_eqvt_at theorem with the indction hypothesis
Christian Urban <urbanc@in.tum.de>
parents: 2654
diff changeset
   609
thm fundef_ex1_eqvt_at
2651
4aa72a88b2c1 equivariance of THE_default under the uniqueness assumption
Christian Urban <urbanc@in.tum.de>
parents: 2650
diff changeset
   610
2649
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   611
ML {*
2655
1c3ad1256f16 instantiated fundef_ex1_eqvt_at theorem with the indction hypothesis
Christian Urban <urbanc@in.tum.de>
parents: 2654
diff changeset
   612
fun prove_stuff ctxt globals G f R clauses complete compat compat_store G_elim G_eqvt f_def =
2649
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   613
  let
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   614
    val Globals {h, domT, ranT, x, ...} = globals
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   615
    val thy = ProofContext.theory_of ctxt
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   616
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   617
    (* Inductive Hypothesis: !!z. (z,x):R ==> EX!y. (z,y):G *)
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   618
    val ihyp = Term.all domT $ Abs ("z", domT,
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   619
      Logic.mk_implies (HOLogic.mk_Trueprop (R $ Bound 0 $ x),
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   620
        HOLogic.mk_Trueprop (Const (@{const_name Ex1}, (ranT --> boolT) --> boolT) $
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   621
          Abs ("y", ranT, G $ Bound 1 $ Bound 0))))
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   622
      |> cterm_of thy
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   623
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   624
    val ihyp_thm = Thm.assume ihyp |> Thm.forall_elim_vars 0
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   625
    val ih_intro = ihyp_thm RS (f_def RS ex1_implies_ex)
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   626
    val ih_elim = ihyp_thm RS (f_def RS ex1_implies_un)
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   627
      |> instantiate' [] [NONE, SOME (cterm_of thy h)]
2655
1c3ad1256f16 instantiated fundef_ex1_eqvt_at theorem with the indction hypothesis
Christian Urban <urbanc@in.tum.de>
parents: 2654
diff changeset
   628
    val ih_eqvt = ihyp_thm RS (G_eqvt RS (f_def RS @{thm fundef_ex1_eqvt_at}))
2649
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   629
2660
3342a2d13d95 nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
parents: 2657
diff changeset
   630
    (*
3342a2d13d95 nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
parents: 2657
diff changeset
   631
    val _ = tracing ("ihyp_thm\n" ^ pp_thm ihyp_thm)
3342a2d13d95 nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
parents: 2657
diff changeset
   632
    val _ = tracing ("ih_intro\n" ^ pp_thm ih_intro)
3342a2d13d95 nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
parents: 2657
diff changeset
   633
    val _ = tracing ("ih_elim\n" ^ pp_thm ih_elim)
3342a2d13d95 nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
parents: 2657
diff changeset
   634
    val _ = tracing ("ih_eqvt\n" ^ pp_thm ih_eqvt)
3342a2d13d95 nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
parents: 2657
diff changeset
   635
    *)
2651
4aa72a88b2c1 equivariance of THE_default under the uniqueness assumption
Christian Urban <urbanc@in.tum.de>
parents: 2650
diff changeset
   636
2649
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   637
    val _ = trace_msg (K "Proving Replacement lemmas...")
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   638
    val repLemmas = map (mk_replacement_lemma thy h ih_elim) clauses
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   639
2660
3342a2d13d95 nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
parents: 2657
diff changeset
   640
    val _ = trace_msg (K "Proving Equivariance lemmas...")
3342a2d13d95 nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
parents: 2657
diff changeset
   641
    val eqvtLemmas = map (mk_eqvt_lemma thy ih_eqvt) clauses
2651
4aa72a88b2c1 equivariance of THE_default under the uniqueness assumption
Christian Urban <urbanc@in.tum.de>
parents: 2650
diff changeset
   642
2649
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   643
    val _ = trace_msg (K "Proving cases for unique existence...")
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   644
    val (ex1s, values) =
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   645
      split_list (map (mk_uniqueness_case thy globals G f 
2660
3342a2d13d95 nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
parents: 2657
diff changeset
   646
        ihyp ih_intro G_elim compat_store clauses repLemmas eqvtLemmas) clauses)
2653
d0f774d06e6e added eqvt_at premises in function definition - however not proved at the moment
Christian Urban <urbanc@in.tum.de>
parents: 2652
diff changeset
   647
     
2649
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   648
    val _ = trace_msg (K "Proving: Graph is a function")
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   649
    val graph_is_function = complete
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   650
      |> Thm.forall_elim_vars 0
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   651
      |> fold (curry op COMP) ex1s
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   652
      |> Thm.implies_intr (ihyp)
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   653
      |> Thm.implies_intr (cterm_of thy (HOLogic.mk_Trueprop (mk_acc domT R $ x)))
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   654
      |> Thm.forall_intr (cterm_of thy x)
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   655
      |> (fn it => Drule.compose_single (it, 2, acc_induct_rule)) (* "EX! y. (?x,y):G" *)
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   656
      |> (fn it => fold (Thm.forall_intr o cterm_of thy o Var) (Term.add_vars (prop_of it) []) it)
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   657
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   658
    val goalstate =  Conjunction.intr graph_is_function complete
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   659
      |> Thm.close_derivation
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   660
      |> Goal.protect
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   661
      |> fold_rev (Thm.implies_intr o cprop_of) compat
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   662
      |> Thm.implies_intr (cprop_of complete)
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   663
  in
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   664
    (goalstate, values)
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   665
  end
2650
e5fa8de0e4bd derived equivariance for the function graph and function relation
Christian Urban <urbanc@in.tum.de>
parents: 2649
diff changeset
   666
*}
2649
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   667
2651
4aa72a88b2c1 equivariance of THE_default under the uniqueness assumption
Christian Urban <urbanc@in.tum.de>
parents: 2650
diff changeset
   668
2650
e5fa8de0e4bd derived equivariance for the function graph and function relation
Christian Urban <urbanc@in.tum.de>
parents: 2649
diff changeset
   669
e5fa8de0e4bd derived equivariance for the function graph and function relation
Christian Urban <urbanc@in.tum.de>
parents: 2649
diff changeset
   670
ML {*
2649
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   671
(* wrapper -- restores quantifiers in rule specifications *)
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   672
fun inductive_def (binding as ((R, T), _)) intrs lthy =
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   673
  let
2650
e5fa8de0e4bd derived equivariance for the function graph and function relation
Christian Urban <urbanc@in.tum.de>
parents: 2649
diff changeset
   674
    val ({intrs = intrs_gen, elims = [elim_gen], preds = [ Rdef ], induct, raw_induct, ...}, lthy) =
2649
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   675
      lthy
2650
e5fa8de0e4bd derived equivariance for the function graph and function relation
Christian Urban <urbanc@in.tum.de>
parents: 2649
diff changeset
   676
      |> Local_Theory.conceal 
2649
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   677
      |> Inductive.add_inductive_i
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   678
          {quiet_mode = true,
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   679
            verbose = ! trace,
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   680
            alt_name = Binding.empty,
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   681
            coind = false,
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   682
            no_elim = false,
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   683
            no_ind = false,
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   684
            skip_mono = true,
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   685
            fork_mono = false}
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   686
          [binding] (* relation *)
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   687
          [] (* no parameters *)
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   688
          (map (fn t => (Attrib.empty_binding, t)) intrs) (* intro rules *)
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   689
          [] (* no special monos *)
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   690
      ||> Local_Theory.restore_naming lthy
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   691
2650
e5fa8de0e4bd derived equivariance for the function graph and function relation
Christian Urban <urbanc@in.tum.de>
parents: 2649
diff changeset
   692
    val ([eqvt_thm], lthy) = Nominal_Eqvt.raw_equivariance false [Rdef] raw_induct intrs_gen lthy
e5fa8de0e4bd derived equivariance for the function graph and function relation
Christian Urban <urbanc@in.tum.de>
parents: 2649
diff changeset
   693
    val eqvt_thm' = (Nominal_ThmDecls.eqvt_transform lthy eqvt_thm) RS @{thm eqvtI}
e5fa8de0e4bd derived equivariance for the function graph and function relation
Christian Urban <urbanc@in.tum.de>
parents: 2649
diff changeset
   694
2649
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   695
    val cert = cterm_of (ProofContext.theory_of lthy)
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   696
    fun requantify orig_intro thm =
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   697
      let
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   698
        val (qs, t) = dest_all_all orig_intro
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   699
        val frees = frees_in_term lthy t |> remove (op =) (Binding.name_of R, T)
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   700
        val vars = Term.add_vars (prop_of thm) [] |> rev
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   701
        val varmap = AList.lookup (op =) (frees ~~ map fst vars)
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   702
          #> the_default ("",0)
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   703
      in
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   704
        fold_rev (fn Free (n, T) =>
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   705
          forall_intr_rename (n, cert (Var (varmap (n, T), T)))) qs thm
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   706
      end
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   707
  in
2650
e5fa8de0e4bd derived equivariance for the function graph and function relation
Christian Urban <urbanc@in.tum.de>
parents: 2649
diff changeset
   708
    ((Rdef, map2 requantify intrs intrs_gen, forall_intr_vars elim_gen, induct, eqvt_thm'), lthy)
2649
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   709
  end
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   710
*}
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   711
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   712
ML {*
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   713
fun define_graph Gname fvar domT ranT clauses RCss lthy =
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   714
  let
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   715
    val GT = domT --> ranT --> boolT
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   716
    val (Gvar as (n, T)) = singleton (Variable.variant_frees lthy []) (Gname, GT)
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   717
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   718
    fun mk_GIntro (ClauseContext {qs, gs, lhs, rhs, ...}) RCs =
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   719
      let
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   720
        fun mk_h_assm (rcfix, rcassm, rcarg) =
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   721
          HOLogic.mk_Trueprop (Free Gvar $ rcarg $ (fvar $ rcarg))
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   722
          |> fold_rev (curry Logic.mk_implies o prop_of) rcassm
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   723
          |> fold_rev (Logic.all o Free) rcfix
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   724
      in
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   725
        HOLogic.mk_Trueprop (Free Gvar $ lhs $ rhs)
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   726
        |> fold_rev (curry Logic.mk_implies o mk_h_assm) RCs
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   727
        |> fold_rev (curry Logic.mk_implies) gs
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   728
        |> fold_rev Logic.all (fvar :: qs)
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   729
      end
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   730
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   731
    val G_intros = map2 mk_GIntro clauses RCss
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   732
  in
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   733
    inductive_def ((Binding.name n, T), NoSyn) G_intros lthy
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   734
  end
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   735
*}
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   736
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   737
ML {*
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   738
fun define_function fdefname (fname, mixfix) domT ranT G default lthy =
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   739
  let
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   740
    val f_def =
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   741
      Abs ("x", domT, Const (@{const_name FunDef.THE_default}, ranT --> (ranT --> boolT) --> ranT) 
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   742
        $ (default $ Bound 0) $ Abs ("y", ranT, G $ Bound 1 $ Bound 0))
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   743
      |> Syntax.check_term lthy
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   744
  in
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   745
    Local_Theory.define
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   746
      ((Binding.name (function_name fname), mixfix),
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   747
        ((Binding.conceal (Binding.name fdefname), []), f_def)) lthy
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   748
  end
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   749
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   750
fun define_recursion_relation Rname domT qglrs clauses RCss lthy =
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   751
  let
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   752
    val RT = domT --> domT --> boolT
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   753
    val (Rvar as (n, T)) = singleton (Variable.variant_frees lthy []) (Rname, RT)
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   754
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   755
    fun mk_RIntro (ClauseContext {qs, gs, lhs, ...}, (oqs, _, _, _)) (rcfix, rcassm, rcarg) =
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   756
      HOLogic.mk_Trueprop (Free Rvar $ rcarg $ lhs)
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   757
      |> fold_rev (curry Logic.mk_implies o prop_of) rcassm
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   758
      |> fold_rev (curry Logic.mk_implies) gs
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   759
      |> fold_rev (Logic.all o Free) rcfix
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   760
      |> fold_rev mk_forall_rename (map fst oqs ~~ qs)
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   761
      (* "!!qs xs. CS ==> G => (r, lhs) : R" *)
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   762
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   763
    val R_intross = map2 (map o mk_RIntro) (clauses ~~ qglrs) RCss
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   764
2650
e5fa8de0e4bd derived equivariance for the function graph and function relation
Christian Urban <urbanc@in.tum.de>
parents: 2649
diff changeset
   765
    val ((R, RIntro_thms, R_elim, _, R_eqvt), lthy) =
2649
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   766
      inductive_def ((Binding.name n, T), NoSyn) (flat R_intross) lthy
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   767
  in
2650
e5fa8de0e4bd derived equivariance for the function graph and function relation
Christian Urban <urbanc@in.tum.de>
parents: 2649
diff changeset
   768
    ((R, Library.unflat R_intross RIntro_thms, R_elim, R_eqvt), lthy)
2649
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   769
  end
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   770
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   771
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   772
fun fix_globals domT ranT fvar ctxt =
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   773
  let
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   774
    val ([h, y, x, z, a, D, P, Pbool],ctxt') = Variable.variant_fixes
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   775
      ["h_fd", "y_fd", "x_fd", "z_fd", "a_fd", "D_fd", "P_fd", "Pb_fd"] ctxt
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   776
  in
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   777
    (Globals {h = Free (h, domT --> ranT),
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   778
      y = Free (y, ranT),
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   779
      x = Free (x, domT),
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   780
      z = Free (z, domT),
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   781
      a = Free (a, domT),
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   782
      D = Free (D, domT --> boolT),
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   783
      P = Free (P, domT --> boolT),
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   784
      Pbool = Free (Pbool, boolT),
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   785
      fvar = fvar,
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   786
      domT = domT,
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   787
      ranT = ranT},
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   788
    ctxt')
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   789
  end
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   790
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   791
fun inst_RC thy fvar f (rcfix, rcassm, rcarg) =
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   792
  let
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   793
    fun inst_term t = subst_bound(f, abstract_over (fvar, t))
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   794
  in
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   795
    (rcfix, map (Thm.assume o cterm_of thy o inst_term o prop_of) rcassm, inst_term rcarg)
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   796
  end
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   797
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   798
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   799
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   800
(**********************************************************
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   801
 *                   PROVING THE RULES
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   802
 **********************************************************)
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   803
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   804
fun mk_psimps thy globals R clauses valthms f_iff graph_is_function =
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   805
  let
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   806
    val Globals {domT, z, ...} = globals
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   807
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   808
    fun mk_psimp (ClauseInfo {qglr = (oqs, _, _, _), cdata = ClauseContext {cqs, lhs, ags, ...}, ...}) valthm =
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   809
      let
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   810
        val lhs_acc = cterm_of thy (HOLogic.mk_Trueprop (mk_acc domT R $ lhs)) (* "acc R lhs" *)
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   811
        val z_smaller = cterm_of thy (HOLogic.mk_Trueprop (R $ z $ lhs)) (* "R z lhs" *)
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   812
      in
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   813
        ((Thm.assume z_smaller) RS ((Thm.assume lhs_acc) RS acc_downward))
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   814
        |> (fn it => it COMP graph_is_function)
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   815
        |> Thm.implies_intr z_smaller
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   816
        |> Thm.forall_intr (cterm_of thy z)
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   817
        |> (fn it => it COMP valthm)
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   818
        |> Thm.implies_intr lhs_acc
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   819
        |> asm_simplify (HOL_basic_ss addsimps [f_iff])
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   820
        |> fold_rev (Thm.implies_intr o cprop_of) ags
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   821
        |> fold_rev forall_intr_rename (map fst oqs ~~ cqs)
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   822
      end
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   823
  in
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   824
    map2 mk_psimp clauses valthms
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   825
  end
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   826
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   827
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   828
(** Induction rule **)
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   829
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   830
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   831
val acc_subset_induct = @{thm predicate1I} RS @{thm accp_subset_induct}
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   832
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   833
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   834
fun mk_partial_induct_rule thy globals R complete_thm clauses =
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   835
  let
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   836
    val Globals {domT, x, z, a, P, D, ...} = globals
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   837
    val acc_R = mk_acc domT R
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   838
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   839
    val x_D = Thm.assume (cterm_of thy (HOLogic.mk_Trueprop (D $ x)))
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   840
    val a_D = cterm_of thy (HOLogic.mk_Trueprop (D $ a))
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   841
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   842
    val D_subset = cterm_of thy (Logic.all x
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   843
      (Logic.mk_implies (HOLogic.mk_Trueprop (D $ x), HOLogic.mk_Trueprop (acc_R $ x))))
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   844
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   845
    val D_dcl = (* "!!x z. [| x: D; (z,x):R |] ==> z:D" *)
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   846
      Logic.all x (Logic.all z (Logic.mk_implies (HOLogic.mk_Trueprop (D $ x),
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   847
        Logic.mk_implies (HOLogic.mk_Trueprop (R $ z $ x),
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   848
          HOLogic.mk_Trueprop (D $ z)))))
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   849
      |> cterm_of thy
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   850
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   851
    (* Inductive Hypothesis: !!z. (z,x):R ==> P z *)
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   852
    val ihyp = Term.all domT $ Abs ("z", domT,
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   853
      Logic.mk_implies (HOLogic.mk_Trueprop (R $ Bound 0 $ x),
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   854
        HOLogic.mk_Trueprop (P $ Bound 0)))
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   855
      |> cterm_of thy
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   856
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   857
    val aihyp = Thm.assume ihyp
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   858
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   859
    fun prove_case clause =
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   860
      let
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   861
        val ClauseInfo {cdata = ClauseContext {ctxt, qs, cqs, ags, gs, lhs, case_hyp, ...},
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   862
          RCs, qglr = (oqs, _, _, _), ...} = clause
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   863
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   864
        val case_hyp_conv = K (case_hyp RS eq_reflection)
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   865
        local open Conv in
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   866
          val lhs_D = fconv_rule (arg_conv (arg_conv (case_hyp_conv))) x_D
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   867
          val sih =
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   868
            fconv_rule (Conv.binder_conv
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   869
              (K (arg1_conv (arg_conv (arg_conv case_hyp_conv)))) ctxt) aihyp
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   870
        end
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   871
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   872
        fun mk_Prec (RCInfo {llRI, RIvs, CCas, rcarg, ...}) = sih
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   873
          |> Thm.forall_elim (cterm_of thy rcarg)
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   874
          |> Thm.elim_implies llRI
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   875
          |> fold_rev (Thm.implies_intr o cprop_of) CCas
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   876
          |> fold_rev (Thm.forall_intr o cterm_of thy o Free) RIvs
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   877
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   878
        val P_recs = map mk_Prec RCs   (*  [P rec1, P rec2, ... ]  *)
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   879
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   880
        val step = HOLogic.mk_Trueprop (P $ lhs)
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   881
          |> fold_rev (curry Logic.mk_implies o prop_of) P_recs
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   882
          |> fold_rev (curry Logic.mk_implies) gs
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   883
          |> curry Logic.mk_implies (HOLogic.mk_Trueprop (D $ lhs))
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   884
          |> fold_rev mk_forall_rename (map fst oqs ~~ qs)
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   885
          |> cterm_of thy
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   886
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   887
        val P_lhs = Thm.assume step
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   888
          |> fold Thm.forall_elim cqs
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   889
          |> Thm.elim_implies lhs_D
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   890
          |> fold Thm.elim_implies ags
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   891
          |> fold Thm.elim_implies P_recs
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   892
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   893
        val res = cterm_of thy (HOLogic.mk_Trueprop (P $ x))
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   894
          |> Conv.arg_conv (Conv.arg_conv case_hyp_conv)
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   895
          |> Thm.symmetric (* P lhs == P x *)
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   896
          |> (fn eql => Thm.equal_elim eql P_lhs) (* "P x" *)
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   897
          |> Thm.implies_intr (cprop_of case_hyp)
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   898
          |> fold_rev (Thm.implies_intr o cprop_of) ags
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   899
          |> fold_rev Thm.forall_intr cqs
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   900
      in
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   901
        (res, step)
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   902
      end
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   903
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   904
    val (cases, steps) = split_list (map prove_case clauses)
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   905
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   906
    val istep = complete_thm
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   907
      |> Thm.forall_elim_vars 0
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   908
      |> fold (curry op COMP) cases (*  P x  *)
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   909
      |> Thm.implies_intr ihyp
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   910
      |> Thm.implies_intr (cprop_of x_D)
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   911
      |> Thm.forall_intr (cterm_of thy x)
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   912
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   913
    val subset_induct_rule =
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   914
      acc_subset_induct
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   915
      |> (curry op COMP) (Thm.assume D_subset)
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   916
      |> (curry op COMP) (Thm.assume D_dcl)
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   917
      |> (curry op COMP) (Thm.assume a_D)
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   918
      |> (curry op COMP) istep
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   919
      |> fold_rev Thm.implies_intr steps
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   920
      |> Thm.implies_intr a_D
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   921
      |> Thm.implies_intr D_dcl
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   922
      |> Thm.implies_intr D_subset
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   923
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   924
    val simple_induct_rule =
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   925
      subset_induct_rule
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   926
      |> Thm.forall_intr (cterm_of thy D)
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   927
      |> Thm.forall_elim (cterm_of thy acc_R)
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   928
      |> assume_tac 1 |> Seq.hd
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   929
      |> (curry op COMP) (acc_downward
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   930
        |> (instantiate' [SOME (ctyp_of thy domT)]
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   931
             (map (SOME o cterm_of thy) [R, x, z]))
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   932
        |> Thm.forall_intr (cterm_of thy z)
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   933
        |> Thm.forall_intr (cterm_of thy x))
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   934
      |> Thm.forall_intr (cterm_of thy a)
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   935
      |> Thm.forall_intr (cterm_of thy P)
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   936
  in
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   937
    simple_induct_rule
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   938
  end
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   939
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   940
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   941
(* FIXME: broken by design *)
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   942
fun mk_domain_intro ctxt (Globals {domT, ...}) R R_cases clause =
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   943
  let
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   944
    val thy = ProofContext.theory_of ctxt
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   945
    val ClauseInfo {cdata = ClauseContext {gs, lhs, cqs, ...},
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   946
      qglr = (oqs, _, _, _), ...} = clause
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   947
    val goal = HOLogic.mk_Trueprop (mk_acc domT R $ lhs)
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   948
      |> fold_rev (curry Logic.mk_implies) gs
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   949
      |> cterm_of thy
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   950
  in
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   951
    Goal.init goal
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   952
    |> (SINGLE (resolve_tac [accI] 1)) |> the
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   953
    |> (SINGLE (eresolve_tac [Thm.forall_elim_vars 0 R_cases] 1))  |> the
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   954
    |> (SINGLE (auto_tac (clasimpset_of ctxt))) |> the
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   955
    |> Goal.conclude
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   956
    |> fold_rev forall_intr_rename (map fst oqs ~~ cqs)
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   957
  end
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   958
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   959
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   960
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   961
(** Termination rule **)
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   962
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   963
val wf_induct_rule = @{thm Wellfounded.wfP_induct_rule}
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   964
val wf_in_rel = @{thm FunDef.wf_in_rel}
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   965
val in_rel_def = @{thm FunDef.in_rel_def}
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   966
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   967
fun mk_nest_term_case thy globals R' ihyp clause =
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   968
  let
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   969
    val Globals {z, ...} = globals
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   970
    val ClauseInfo {cdata = ClauseContext {qs, cqs, ags, lhs, case_hyp, ...}, tree,
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   971
      qglr=(oqs, _, _, _), ...} = clause
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   972
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   973
    val ih_case = full_simplify (HOL_basic_ss addsimps [case_hyp]) ihyp
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   974
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   975
    fun step (fixes, assumes) (_ $ arg) u (sub,(hyps,thms)) =
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   976
      let
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   977
        val used = (u @ sub)
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   978
          |> map (fn (ctx,thm) => Function_Ctx_Tree.export_thm thy ctx thm)
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   979
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   980
        val hyp = HOLogic.mk_Trueprop (R' $ arg $ lhs)
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   981
          |> fold_rev (curry Logic.mk_implies o prop_of) used (* additional hyps *)
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   982
          |> Function_Ctx_Tree.export_term (fixes, assumes)
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   983
          |> fold_rev (curry Logic.mk_implies o prop_of) ags
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   984
          |> fold_rev mk_forall_rename (map fst oqs ~~ qs)
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   985
          |> cterm_of thy
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   986
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   987
        val thm = Thm.assume hyp
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   988
          |> fold Thm.forall_elim cqs
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   989
          |> fold Thm.elim_implies ags
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   990
          |> Function_Ctx_Tree.import_thm thy (fixes, assumes)
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   991
          |> fold Thm.elim_implies used (*  "(arg, lhs) : R'"  *)
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   992
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   993
        val z_eq_arg = HOLogic.mk_Trueprop (mk_eq (z, arg))
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   994
          |> cterm_of thy |> Thm.assume
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   995
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   996
        val acc = thm COMP ih_case
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   997
        val z_acc_local = acc
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   998
          |> Conv.fconv_rule
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   999
              (Conv.arg_conv (Conv.arg_conv (K (Thm.symmetric (z_eq_arg RS eq_reflection)))))
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1000
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1001
        val ethm = z_acc_local
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1002
          |> Function_Ctx_Tree.export_thm thy (fixes,
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1003
               z_eq_arg :: case_hyp :: ags @ assumes)
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1004
          |> fold_rev forall_intr_rename (map fst oqs ~~ cqs)
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1005
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1006
        val sub' = sub @ [(([],[]), acc)]
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1007
      in
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1008
        (sub', (hyp :: hyps, ethm :: thms))
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1009
      end
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1010
      | step _ _ _ _ = raise Match
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1011
  in
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1012
    Function_Ctx_Tree.traverse_tree step tree
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1013
  end
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1014
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1015
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1016
fun mk_nest_term_rule thy globals R R_cases clauses =
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1017
  let
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1018
    val Globals { domT, x, z, ... } = globals
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1019
    val acc_R = mk_acc domT R
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1020
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1021
    val R' = Free ("R", fastype_of R)
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1022
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1023
    val Rrel = Free ("R", HOLogic.mk_setT (HOLogic.mk_prodT (domT, domT)))
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1024
    val inrel_R = Const (@{const_name FunDef.in_rel},
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1025
      HOLogic.mk_setT (HOLogic.mk_prodT (domT, domT)) --> fastype_of R) $ Rrel
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1026
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1027
    val wfR' = HOLogic.mk_Trueprop (Const (@{const_name Wellfounded.wfP},
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1028
      (domT --> domT --> boolT) --> boolT) $ R')
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1029
      |> cterm_of thy (* "wf R'" *)
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1030
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1031
    (* Inductive Hypothesis: !!z. (z,x):R' ==> z : acc R *)
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1032
    val ihyp = Term.all domT $ Abs ("z", domT,
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1033
      Logic.mk_implies (HOLogic.mk_Trueprop (R' $ Bound 0 $ x),
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1034
        HOLogic.mk_Trueprop (acc_R $ Bound 0)))
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1035
      |> cterm_of thy
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1036
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1037
    val ihyp_a = Thm.assume ihyp |> Thm.forall_elim_vars 0
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1038
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1039
    val R_z_x = cterm_of thy (HOLogic.mk_Trueprop (R $ z $ x))
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1040
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1041
    val (hyps, cases) = fold (mk_nest_term_case thy globals R' ihyp_a) clauses ([], [])
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1042
  in
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1043
    R_cases
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1044
    |> Thm.forall_elim (cterm_of thy z)
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1045
    |> Thm.forall_elim (cterm_of thy x)
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1046
    |> Thm.forall_elim (cterm_of thy (acc_R $ z))
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1047
    |> curry op COMP (Thm.assume R_z_x)
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1048
    |> fold_rev (curry op COMP) cases
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1049
    |> Thm.implies_intr R_z_x
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1050
    |> Thm.forall_intr (cterm_of thy z)
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1051
    |> (fn it => it COMP accI)
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1052
    |> Thm.implies_intr ihyp
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1053
    |> Thm.forall_intr (cterm_of thy x)
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1054
    |> (fn it => Drule.compose_single(it,2,wf_induct_rule))
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1055
    |> curry op RS (Thm.assume wfR')
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1056
    |> forall_intr_vars
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1057
    |> (fn it => it COMP allI)
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1058
    |> fold Thm.implies_intr hyps
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1059
    |> Thm.implies_intr wfR'
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1060
    |> Thm.forall_intr (cterm_of thy R')
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1061
    |> Thm.forall_elim (cterm_of thy (inrel_R))
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1062
    |> curry op RS wf_in_rel
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1063
    |> full_simplify (HOL_basic_ss addsimps [in_rel_def])
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1064
    |> Thm.forall_intr (cterm_of thy Rrel)
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1065
  end
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1066
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1067
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1068
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1069
(* Tail recursion (probably very fragile)
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1070
 *
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1071
 * FIXME:
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1072
 * - Need to do forall_elim_vars on psimps: Unneccesary, if psimps would be taken from the same context.
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1073
 * - Must we really replace the fvar by f here?
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1074
 * - Splitting is not configured automatically: Problems with case?
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1075
 *)
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1076
fun mk_trsimps octxt globals f G R f_def R_cases G_induct clauses psimps =
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1077
  let
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1078
    val Globals {domT, ranT, fvar, ...} = globals
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1079
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1080
    val R_cases = Thm.forall_elim_vars 0 R_cases (* FIXME: Should be already in standard form. *)
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1081
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1082
    val graph_implies_dom = (* "G ?x ?y ==> dom ?x"  *)
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1083
      Goal.prove octxt ["x", "y"] [HOLogic.mk_Trueprop (G $ Free ("x", domT) $ Free ("y", ranT))]
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1084
        (HOLogic.mk_Trueprop (mk_acc domT R $ Free ("x", domT)))
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1085
        (fn {prems=[a], ...} =>
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1086
          ((rtac (G_induct OF [a]))
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1087
          THEN_ALL_NEW rtac accI
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1088
          THEN_ALL_NEW etac R_cases
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1089
          THEN_ALL_NEW asm_full_simp_tac (simpset_of octxt)) 1)
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1090
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1091
    val default_thm =
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1092
      forall_intr_vars graph_implies_dom COMP (f_def COMP fundef_default_value)
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1093
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1094
    fun mk_trsimp clause psimp =
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1095
      let
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1096
        val ClauseInfo {qglr = (oqs, _, _, _), cdata =
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1097
          ClauseContext {ctxt, cqs, gs, lhs, rhs, ...}, ...} = clause
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1098
        val thy = ProofContext.theory_of ctxt
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1099
        val rhs_f = Pattern.rewrite_term thy [(fvar, f)] [] rhs
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1100
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1101
        val trsimp = Logic.list_implies(gs,
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1102
          HOLogic.mk_Trueprop (HOLogic.mk_eq(f $ lhs, rhs_f))) (* "f lhs = rhs" *)
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1103
        val lhs_acc = (mk_acc domT R $ lhs) (* "acc R lhs" *)
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1104
        fun simp_default_tac ss =
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1105
          asm_full_simp_tac (ss addsimps [default_thm, Let_def])
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1106
      in
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1107
        Goal.prove ctxt [] [] trsimp (fn _ =>
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1108
          rtac (instantiate' [] [SOME (cterm_of thy lhs_acc)] case_split) 1
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1109
          THEN (rtac (Thm.forall_elim_vars 0 psimp) THEN_ALL_NEW assume_tac) 1
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1110
          THEN (simp_default_tac (simpset_of ctxt) 1)
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1111
          THEN TRY ((etac not_acc_down 1)
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1112
            THEN ((etac R_cases)
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1113
              THEN_ALL_NEW (simp_default_tac (simpset_of ctxt))) 1))
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1114
        |> fold_rev forall_intr_rename (map fst oqs ~~ cqs)
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1115
      end
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1116
  in
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1117
    map2 mk_trsimp clauses psimps
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1118
  end
2650
e5fa8de0e4bd derived equivariance for the function graph and function relation
Christian Urban <urbanc@in.tum.de>
parents: 2649
diff changeset
  1119
*}
2649
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1120
2650
e5fa8de0e4bd derived equivariance for the function graph and function relation
Christian Urban <urbanc@in.tum.de>
parents: 2649
diff changeset
  1121
ML {*
2649
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1122
fun prepare_function config defname [((fname, fT), mixfix)] abstract_qglrs lthy =
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1123
  let
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1124
    val FunctionConfig {domintros, tailrec, default=default_opt, ...} = config
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1125
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1126
    val default_str = the_default "%x. undefined" default_opt (*FIXME dynamic scoping*)
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1127
    val fvar = Free (fname, fT)
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1128
    val domT = domain_type fT
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1129
    val ranT = range_type fT
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1130
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1131
    val default = Syntax.parse_term lthy default_str
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1132
      |> Type.constraint fT |> Syntax.check_term lthy
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1133
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1134
    val (globals, ctxt') = fix_globals domT ranT fvar lthy
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1135
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1136
    val Globals { x, h, ... } = globals
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1137
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1138
    val clauses = map (mk_clause_context x ctxt') abstract_qglrs
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1139
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1140
    val n = length abstract_qglrs
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1141
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1142
    fun build_tree (ClauseContext { ctxt, rhs, ...}) =
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1143
       Function_Ctx_Tree.mk_tree (fname, fT) h ctxt rhs
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1144
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1145
    val trees = map build_tree clauses
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1146
    val RCss = map find_calls trees
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1147
2650
e5fa8de0e4bd derived equivariance for the function graph and function relation
Christian Urban <urbanc@in.tum.de>
parents: 2649
diff changeset
  1148
    val ((G, GIntro_thms, G_elim, G_induct, G_eqvt), lthy) =
2649
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1149
      PROFILE "def_graph" (define_graph (graph_name defname) fvar domT ranT clauses RCss) lthy
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1150
2653
d0f774d06e6e added eqvt_at premises in function definition - however not proved at the moment
Christian Urban <urbanc@in.tum.de>
parents: 2652
diff changeset
  1151
    (* 
2660
3342a2d13d95 nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
parents: 2657
diff changeset
  1152
    val _ = tracing ("Graph - name: " ^ pp_thm G)
3342a2d13d95 nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
parents: 2657
diff changeset
  1153
    val _ = tracing ("Graph intros:\n" ^ cat_lines (map pp_thm GIntro_thms))
3342a2d13d95 nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
parents: 2657
diff changeset
  1154
    val _ = tracing ("Graph Equivariance " ^ pp_thm G_eqvt)
2653
d0f774d06e6e added eqvt_at premises in function definition - however not proved at the moment
Christian Urban <urbanc@in.tum.de>
parents: 2652
diff changeset
  1155
    *)
2649
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1156
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1157
    val ((f, (_, f_defthm)), lthy) =
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1158
      PROFILE "def_fun" (define_function (defname ^ "_sumC_def") (fname, mixfix) domT ranT G default) lthy
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1159
2653
d0f774d06e6e added eqvt_at premises in function definition - however not proved at the moment
Christian Urban <urbanc@in.tum.de>
parents: 2652
diff changeset
  1160
    (*
2660
3342a2d13d95 nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
parents: 2657
diff changeset
  1161
    val _ = tracing ("f - name: " ^ pp_thm f)
3342a2d13d95 nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
parents: 2657
diff changeset
  1162
    val _ = tracing ("f_defthm:\n" ^ pp_thm f_defthm)
2653
d0f774d06e6e added eqvt_at premises in function definition - however not proved at the moment
Christian Urban <urbanc@in.tum.de>
parents: 2652
diff changeset
  1163
    *)
2649
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1164
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1165
    val RCss = map (map (inst_RC (ProofContext.theory_of lthy) fvar f)) RCss
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1166
    val trees = map (Function_Ctx_Tree.inst_tree (ProofContext.theory_of lthy) fvar f) trees
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1167
2655
1c3ad1256f16 instantiated fundef_ex1_eqvt_at theorem with the indction hypothesis
Christian Urban <urbanc@in.tum.de>
parents: 2654
diff changeset
  1168
    (*
2660
3342a2d13d95 nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
parents: 2657
diff changeset
  1169
    val _ = tracing ("recursive calls:\n" ^ cat_lines (map pp_thm RCss))
2655
1c3ad1256f16 instantiated fundef_ex1_eqvt_at theorem with the indction hypothesis
Christian Urban <urbanc@in.tum.de>
parents: 2654
diff changeset
  1170
    *)
2653
d0f774d06e6e added eqvt_at premises in function definition - however not proved at the moment
Christian Urban <urbanc@in.tum.de>
parents: 2652
diff changeset
  1171
2650
e5fa8de0e4bd derived equivariance for the function graph and function relation
Christian Urban <urbanc@in.tum.de>
parents: 2649
diff changeset
  1172
    val ((R, RIntro_thmss, R_elim, R_eqvt), lthy) =
2649
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1173
      PROFILE "def_rel" (define_recursion_relation (rel_name defname) domT abstract_qglrs clauses RCss) lthy
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1174
2653
d0f774d06e6e added eqvt_at premises in function definition - however not proved at the moment
Christian Urban <urbanc@in.tum.de>
parents: 2652
diff changeset
  1175
    (*
2660
3342a2d13d95 nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
parents: 2657
diff changeset
  1176
    val _ = tracing ("Relation - name: " ^ pp_thm R) 
3342a2d13d95 nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
parents: 2657
diff changeset
  1177
    val _ = tracing ("Relation intros:\n" ^ cat_lines (map pp_thm RIntro_thmss))
3342a2d13d95 nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
parents: 2657
diff changeset
  1178
    val _ = tracing ("Relation Equivariance " ^ pp_thm R_eqvt)
2653
d0f774d06e6e added eqvt_at premises in function definition - however not proved at the moment
Christian Urban <urbanc@in.tum.de>
parents: 2652
diff changeset
  1179
    *)    
2649
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1180
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1181
    val (_, lthy) =
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1182
      Local_Theory.abbrev Syntax.mode_default ((Binding.name (dom_name defname), NoSyn), mk_acc domT R) lthy
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1183
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1184
    val newthy = ProofContext.theory_of lthy
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1185
    val clauses = map (transfer_clause_ctx newthy) clauses
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1186
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1187
    val cert = cterm_of (ProofContext.theory_of lthy)
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1188
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1189
    val xclauses = PROFILE "xclauses"
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1190
      (map7 (mk_clause_info globals G f) (1 upto n) clauses abstract_qglrs trees
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1191
        RCss GIntro_thms) RIntro_thmss
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1192
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1193
    val complete =
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1194
      mk_completeness globals clauses abstract_qglrs |> cert |> Thm.assume
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1195
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1196
    val compat =
2653
d0f774d06e6e added eqvt_at premises in function definition - however not proved at the moment
Christian Urban <urbanc@in.tum.de>
parents: 2652
diff changeset
  1197
      mk_compat_proof_obligations domT ranT fvar f abstract_qglrs 
2649
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1198
      |> map (cert #> Thm.assume)
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1199
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1200
    val compat_store = store_compat_thms n compat
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1201
2650
e5fa8de0e4bd derived equivariance for the function graph and function relation
Christian Urban <urbanc@in.tum.de>
parents: 2649
diff changeset
  1202
    (*
2660
3342a2d13d95 nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
parents: 2657
diff changeset
  1203
    val _ = tracing ("globals:\n" ^ pp_thm globals)
3342a2d13d95 nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
parents: 2657
diff changeset
  1204
    val _ = tracing ("complete:\n" ^ pp_thm complete)
3342a2d13d95 nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
parents: 2657
diff changeset
  1205
    val _ = tracing ("compat:\n" ^ pp_thm compat)
3342a2d13d95 nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
parents: 2657
diff changeset
  1206
    val _ = tracing ("compat_store:\n" ^ pp_thm compat_store)
2650
e5fa8de0e4bd derived equivariance for the function graph and function relation
Christian Urban <urbanc@in.tum.de>
parents: 2649
diff changeset
  1207
    *)
2649
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1208
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1209
    val (goalstate, values) = PROFILE "prove_stuff"
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1210
      (prove_stuff lthy globals G f R xclauses complete compat
2655
1c3ad1256f16 instantiated fundef_ex1_eqvt_at theorem with the indction hypothesis
Christian Urban <urbanc@in.tum.de>
parents: 2654
diff changeset
  1211
         compat_store G_elim G_eqvt) f_defthm
2649
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1212
     
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1213
    val mk_trsimps =
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1214
      mk_trsimps lthy globals f G R f_defthm R_elim G_induct xclauses
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1215
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1216
    fun mk_partial_rules provedgoal =
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1217
      let
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1218
        val newthy = theory_of_thm provedgoal (*FIXME*)
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1219
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1220
        val (graph_is_function, complete_thm) =
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1221
          provedgoal
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1222
          |> Conjunction.elim
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1223
          |> apfst (Thm.forall_elim_vars 0)
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1224
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1225
        val f_iff = graph_is_function RS (f_defthm RS ex1_implies_iff)
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1226
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1227
        val psimps = PROFILE "Proving simplification rules"
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1228
          (mk_psimps newthy globals R xclauses values f_iff) graph_is_function
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1229
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1230
        val simple_pinduct = PROFILE "Proving partial induction rule"
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1231
          (mk_partial_induct_rule newthy globals R complete_thm) xclauses
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1232
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1233
        val total_intro = PROFILE "Proving nested termination rule"
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1234
          (mk_nest_term_rule newthy globals R R_elim) xclauses
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1235
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1236
        val dom_intros =
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1237
          if domintros then SOME (PROFILE "Proving domain introduction rules"
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1238
             (map (mk_domain_intro lthy globals R R_elim)) xclauses)
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1239
           else NONE
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1240
        val trsimps = if tailrec then SOME (mk_trsimps psimps) else NONE
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1241
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1242
      in
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1243
        FunctionResult {fs=[f], G=G, R=R, cases=complete_thm,
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1244
          psimps=psimps, simple_pinducts=[simple_pinduct],
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1245
          termination=total_intro, trsimps=trsimps,
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1246
          domintros=dom_intros}
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1247
      end
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1248
  in
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1249
    ((f, goalstate, mk_partial_rules), lthy)
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1250
  end
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1251
*}
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1252
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1253
ML {*
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1254
open Function_Lib
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1255
open Function_Common
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1256
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1257
type qgar = string * (string * typ) list * term list * term list * term
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1258
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1259
datatype mutual_part = MutualPart of
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1260
 {i : int,
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1261
  i' : int,
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1262
  fvar : string * typ,
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1263
  cargTs: typ list,
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1264
  f_def: term,
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1265
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1266
  f: term option,
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1267
  f_defthm : thm option}
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1268
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1269
datatype mutual_info = Mutual of
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1270
 {n : int,
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1271
  n' : int,
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1272
  fsum_var : string * typ,
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1273
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1274
  ST: typ,
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1275
  RST: typ,
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1276
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1277
  parts: mutual_part list,
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1278
  fqgars: qgar list,
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1279
  qglrs: ((string * typ) list * term list * term * term) list,
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1280
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1281
  fsum : term option}
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1282
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1283
fun mutual_induct_Pnames n =
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1284
  if n < 5 then fst (chop n ["P","Q","R","S"])
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1285
  else map (fn i => "P" ^ string_of_int i) (1 upto n)
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1286
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1287
fun get_part fname =
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1288
  the o find_first (fn (MutualPart {fvar=(n,_), ...}) => n = fname)
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1289
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1290
(* FIXME *)
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1291
fun mk_prod_abs e (t1, t2) =
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1292
  let
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1293
    val bTs = rev (map snd e)
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1294
    val T1 = fastype_of1 (bTs, t1)
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1295
    val T2 = fastype_of1 (bTs, t2)
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1296
  in
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1297
    HOLogic.pair_const T1 T2 $ t1 $ t2
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1298
  end
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1299
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1300
fun analyze_eqs ctxt defname fs eqs =
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1301
  let
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1302
    val num = length fs
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1303
    val fqgars = map (split_def ctxt (K true)) eqs
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1304
    val arity_of = map (fn (fname,_,_,args,_) => (fname, length args)) fqgars
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1305
      |> AList.lookup (op =) #> the
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1306
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1307
    fun curried_types (fname, fT) =
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1308
      let
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1309
        val (caTs, uaTs) = chop (arity_of fname) (binder_types fT)
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1310
      in
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1311
        (caTs, uaTs ---> body_type fT)
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1312
      end
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1313
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1314
    val (caTss, resultTs) = split_list (map curried_types fs)
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1315
    val argTs = map (foldr1 HOLogic.mk_prodT) caTss
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1316
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1317
    val dresultTs = distinct (op =) resultTs
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1318
    val n' = length dresultTs
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1319
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1320
    val RST = Balanced_Tree.make (uncurry SumTree.mk_sumT) dresultTs
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1321
    val ST = Balanced_Tree.make (uncurry SumTree.mk_sumT) argTs
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1322
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1323
    val fsum_type = ST --> RST
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1324
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1325
    val ([fsum_var_name], _) = Variable.add_fixes [ defname ^ "_sum" ] ctxt
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1326
    val fsum_var = (fsum_var_name, fsum_type)
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1327
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1328
    fun define (fvar as (n, _)) caTs resultT i =
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1329
      let
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1330
        val vars = map_index (fn (j,T) => Free ("x" ^ string_of_int j, T)) caTs (* FIXME: Bind xs properly *)
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1331
        val i' = find_index (fn Ta => Ta = resultT) dresultTs + 1
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1332
2653
d0f774d06e6e added eqvt_at premises in function definition - however not proved at the moment
Christian Urban <urbanc@in.tum.de>
parents: 2652
diff changeset
  1333
        val f_exp = SumTree.mk_proj RST n' i' 
d0f774d06e6e added eqvt_at premises in function definition - however not proved at the moment
Christian Urban <urbanc@in.tum.de>
parents: 2652
diff changeset
  1334
          (Free fsum_var $ SumTree.mk_inj ST num i (foldr1 HOLogic.mk_prod vars))
d0f774d06e6e added eqvt_at premises in function definition - however not proved at the moment
Christian Urban <urbanc@in.tum.de>
parents: 2652
diff changeset
  1335
        
2649
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1336
        val def = Term.abstract_over (Free fsum_var, fold_rev lambda vars f_exp)
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1337
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1338
        val rew = (n, fold_rev lambda vars f_exp)
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1339
      in
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1340
        (MutualPart {i=i, i'=i', fvar=fvar,cargTs=caTs,f_def=def,f=NONE,f_defthm=NONE}, rew)
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1341
      end
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1342
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1343
    val (parts, rews) = split_list (map4 define fs caTss resultTs (1 upto num))
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1344
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1345
    fun convert_eqs (f, qs, gs, args, rhs) =
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1346
      let
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1347
        val MutualPart {i, i', ...} = get_part f parts
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1348
      in
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1349
        (qs, gs, SumTree.mk_inj ST num i (foldr1 (mk_prod_abs qs) args),
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1350
         SumTree.mk_inj RST n' i' (replace_frees rews rhs)
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1351
         |> Envir.beta_norm)
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1352
      end
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1353
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1354
    val qglrs = map convert_eqs fqgars
2653
d0f774d06e6e added eqvt_at premises in function definition - however not proved at the moment
Christian Urban <urbanc@in.tum.de>
parents: 2652
diff changeset
  1355
   
d0f774d06e6e added eqvt_at premises in function definition - however not proved at the moment
Christian Urban <urbanc@in.tum.de>
parents: 2652
diff changeset
  1356
    fun pp_f (f, args, precond, lhs, rhs) =
d0f774d06e6e added eqvt_at premises in function definition - however not proved at the moment
Christian Urban <urbanc@in.tum.de>
parents: 2652
diff changeset
  1357
      (tracing ("lhs: " ^ commas 
d0f774d06e6e added eqvt_at premises in function definition - however not proved at the moment
Christian Urban <urbanc@in.tum.de>
parents: 2652
diff changeset
  1358
         (map (fn t => Syntax.string_of_term ctxt (subst_bounds (map Free (rev args), t))) lhs));
d0f774d06e6e added eqvt_at premises in function definition - however not proved at the moment
Christian Urban <urbanc@in.tum.de>
parents: 2652
diff changeset
  1359
       tracing ("rhs: " ^ Syntax.string_of_term ctxt (subst_bounds (map Free (rev args), rhs))))
d0f774d06e6e added eqvt_at premises in function definition - however not proved at the moment
Christian Urban <urbanc@in.tum.de>
parents: 2652
diff changeset
  1360
d0f774d06e6e added eqvt_at premises in function definition - however not proved at the moment
Christian Urban <urbanc@in.tum.de>
parents: 2652
diff changeset
  1361
    fun pp_q (args, precond, lhs, rhs) =
d0f774d06e6e added eqvt_at premises in function definition - however not proved at the moment
Christian Urban <urbanc@in.tum.de>
parents: 2652
diff changeset
  1362
      (tracing ("qlhs: " ^ Syntax.string_of_term ctxt (subst_bounds (map Free (rev args), lhs)));
d0f774d06e6e added eqvt_at premises in function definition - however not proved at the moment
Christian Urban <urbanc@in.tum.de>
parents: 2652
diff changeset
  1363
       tracing ("qrhs: " ^ Syntax.string_of_term ctxt (subst_bounds (map Free (rev args), rhs))))
d0f774d06e6e added eqvt_at premises in function definition - however not proved at the moment
Christian Urban <urbanc@in.tum.de>
parents: 2652
diff changeset
  1364
2649
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1365
  in
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1366
    Mutual {n=num, n'=n', fsum_var=fsum_var, ST=ST, RST=RST,
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1367
      parts=parts, fqgars=fqgars, qglrs=qglrs, fsum=NONE}
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1368
  end
2653
d0f774d06e6e added eqvt_at premises in function definition - however not proved at the moment
Christian Urban <urbanc@in.tum.de>
parents: 2652
diff changeset
  1369
*}
2649
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1370
2653
d0f774d06e6e added eqvt_at premises in function definition - however not proved at the moment
Christian Urban <urbanc@in.tum.de>
parents: 2652
diff changeset
  1371
ML {*
2649
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1372
fun define_projections fixes mutual fsum lthy =
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1373
  let
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1374
    fun def ((MutualPart {i=i, i'=i', fvar=(fname, fT), cargTs, f_def, ...}), (_, mixfix)) lthy =
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1375
      let
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1376
        val ((f, (_, f_defthm)), lthy') =
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1377
          Local_Theory.define
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1378
            ((Binding.name fname, mixfix),
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1379
              ((Binding.conceal (Binding.name (fname ^ "_def")), []),
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1380
              Term.subst_bound (fsum, f_def))) lthy
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1381
      in
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1382
        (MutualPart {i=i, i'=i', fvar=(fname, fT), cargTs=cargTs, f_def=f_def,
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1383
           f=SOME f, f_defthm=SOME f_defthm },
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1384
         lthy')
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1385
      end
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1386
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1387
    val Mutual { n, n', fsum_var, ST, RST, parts, fqgars, qglrs, ... } = mutual
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1388
    val (parts', lthy') = fold_map def (parts ~~ fixes) lthy
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1389
  in
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1390
    (Mutual { n=n, n'=n', fsum_var=fsum_var, ST=ST, RST=RST, parts=parts',
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1391
       fqgars=fqgars, qglrs=qglrs, fsum=SOME fsum },
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1392
     lthy')
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1393
  end
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1394
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1395
fun in_context ctxt (f, pre_qs, pre_gs, pre_args, pre_rhs) F =
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1396
  let
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1397
    val thy = ProofContext.theory_of ctxt
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1398
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1399
    val oqnames = map fst pre_qs
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1400
    val (qs, _) = Variable.variant_fixes oqnames ctxt
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1401
      |>> map2 (fn (_, T) => fn n => Free (n, T)) pre_qs
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1402
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1403
    fun inst t = subst_bounds (rev qs, t)
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1404
    val gs = map inst pre_gs
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1405
    val args = map inst pre_args
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1406
    val rhs = inst pre_rhs
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1407
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1408
    val cqs = map (cterm_of thy) qs
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1409
    val ags = map (Thm.assume o cterm_of thy) gs
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1410
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1411
    val import = fold Thm.forall_elim cqs
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1412
      #> fold Thm.elim_implies ags
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1413
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1414
    val export = fold_rev (Thm.implies_intr o cprop_of) ags
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1415
      #> fold_rev forall_intr_rename (oqnames ~~ cqs)
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1416
  in
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1417
    F ctxt (f, qs, gs, args, rhs) import export
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1418
  end
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1419
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1420
fun recover_mutual_psimp all_orig_fdefs parts ctxt (fname, _, _, args, rhs)
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1421
  import (export : thm -> thm) sum_psimp_eq =
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1422
  let
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1423
    val (MutualPart {f=SOME f, ...}) = get_part fname parts
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1424
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1425
    val psimp = import sum_psimp_eq
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1426
    val (simp, restore_cond) =
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1427
      case cprems_of psimp of
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1428
        [] => (psimp, I)
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1429
      | [cond] => (Thm.implies_elim psimp (Thm.assume cond), Thm.implies_intr cond)
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1430
      | _ => raise General.Fail "Too many conditions"
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1431
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1432
  in
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1433
    Goal.prove ctxt [] []
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1434
      (HOLogic.Trueprop $ HOLogic.mk_eq (list_comb (f, args), rhs))
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1435
      (fn _ => (Local_Defs.unfold_tac ctxt all_orig_fdefs)
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1436
         THEN EqSubst.eqsubst_tac ctxt [0] [simp] 1
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1437
         THEN (simp_tac (simpset_of ctxt)) 1) (* FIXME: global simpset?!! *)
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1438
    |> restore_cond
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1439
    |> export
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1440
  end
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1441
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1442
fun mk_applied_form ctxt caTs thm =
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1443
  let
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1444
    val thy = ProofContext.theory_of ctxt
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1445
    val xs = map_index (fn (i,T) => cterm_of thy (Free ("x" ^ string_of_int i, T))) caTs (* FIXME: Bind xs properly *)
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1446
  in
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1447
    fold (fn x => fn thm => Thm.combination thm (Thm.reflexive x)) xs thm
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1448
    |> Conv.fconv_rule (Thm.beta_conversion true)
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1449
    |> fold_rev Thm.forall_intr xs
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1450
    |> Thm.forall_elim_vars 0
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1451
  end
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1452
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1453
fun mutual_induct_rules lthy induct all_f_defs (Mutual {n, ST, parts, ...}) =
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1454
  let
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1455
    val cert = cterm_of (ProofContext.theory_of lthy)
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1456
    val newPs =
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1457
      map2 (fn Pname => fn MutualPart {cargTs, ...} =>
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1458
          Free (Pname, cargTs ---> HOLogic.boolT))
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1459
        (mutual_induct_Pnames (length parts)) parts
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1460
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1461
    fun mk_P (MutualPart {cargTs, ...}) P =
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1462
      let
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1463
        val avars = map_index (fn (i,T) => Var (("a", i), T)) cargTs
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1464
        val atup = foldr1 HOLogic.mk_prod avars
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1465
      in
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1466
        HOLogic.tupled_lambda atup (list_comb (P, avars))
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1467
      end
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1468
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1469
    val Ps = map2 mk_P parts newPs
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1470
    val case_exp = SumTree.mk_sumcases HOLogic.boolT Ps
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1471
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1472
    val induct_inst =
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1473
      Thm.forall_elim (cert case_exp) induct
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1474
      |> full_simplify SumTree.sumcase_split_ss
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1475
      |> full_simplify (HOL_basic_ss addsimps all_f_defs)
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1476
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1477
    fun project rule (MutualPart {cargTs, i, ...}) k =
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1478
      let
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1479
        val afs = map_index (fn (j,T) => Free ("a" ^ string_of_int (j + k), T)) cargTs (* FIXME! *)
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1480
        val inj = SumTree.mk_inj ST n i (foldr1 HOLogic.mk_prod afs)
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1481
      in
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1482
        (rule
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1483
         |> Thm.forall_elim (cert inj)
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1484
         |> full_simplify SumTree.sumcase_split_ss
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1485
         |> fold_rev (Thm.forall_intr o cert) (afs @ newPs),
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1486
         k + length cargTs)
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1487
      end
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1488
  in
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1489
    fst (fold_map (project induct_inst) parts 0)
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1490
  end
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1491
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1492
fun mk_partial_rules_mutual lthy inner_cont (m as Mutual {parts, fqgars, ...}) proof =
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1493
  let
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1494
    val result = inner_cont proof
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1495
    val FunctionResult {G, R, cases, psimps, trsimps, simple_pinducts=[simple_pinduct],
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1496
      termination, domintros, ...} = result
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1497
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1498
    val (all_f_defs, fs) =
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1499
      map (fn MutualPart {f_defthm = SOME f_def, f = SOME f, cargTs, ...} =>
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1500
        (mk_applied_form lthy cargTs (Thm.symmetric f_def), f))
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1501
      parts
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1502
      |> split_list
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1503
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1504
    val all_orig_fdefs =
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1505
      map (fn MutualPart {f_defthm = SOME f_def, ...} => f_def) parts
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1506
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1507
    fun mk_mpsimp fqgar sum_psimp =
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1508
      in_context lthy fqgar (recover_mutual_psimp all_orig_fdefs parts) sum_psimp
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1509
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1510
    val rew_ss = HOL_basic_ss addsimps all_f_defs
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1511
    val mpsimps = map2 mk_mpsimp fqgars psimps
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1512
    val mtrsimps = Option.map (map2 mk_mpsimp fqgars) trsimps
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1513
    val minducts = mutual_induct_rules lthy simple_pinduct all_f_defs m
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1514
    val mtermination = full_simplify rew_ss termination
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1515
    val mdomintros = Option.map (map (full_simplify rew_ss)) domintros
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1516
  in
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1517
    FunctionResult { fs=fs, G=G, R=R,
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1518
      psimps=mpsimps, simple_pinducts=minducts,
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1519
      cases=cases, termination=mtermination,
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1520
      domintros=mdomintros, trsimps=mtrsimps}
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1521
  end
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1522
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1523
fun prepare_function_mutual config defname fixes eqss lthy =
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1524
  let
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1525
    val mutual as Mutual {fsum_var=(n, T), qglrs, ...} =
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1526
      analyze_eqs lthy defname (map fst fixes) (map Envir.beta_eta_contract eqss)
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1527
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1528
    val ((fsum, goalstate, cont), lthy') =
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1529
      prepare_function config defname [((n, T), NoSyn)] qglrs lthy
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1530
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1531
    val (mutual', lthy'') = define_projections fixes mutual fsum lthy'
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1532
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1533
    val mutual_cont = mk_partial_rules_mutual lthy'' cont mutual'
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1534
  in
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1535
    ((goalstate, mutual_cont), lthy'')
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1536
  end
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1537
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1538
*}
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1539
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1540
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1541
ML {*
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1542
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1543
open Function_Lib
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1544
open Function_Common
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1545
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1546
val simp_attribs = map (Attrib.internal o K)
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1547
  [Simplifier.simp_add,
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1548
   Code.add_default_eqn_attribute,
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1549
   Nitpick_Simps.add]
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1550
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1551
val psimp_attribs = map (Attrib.internal o K)
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1552
  [Nitpick_Psimps.add]
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1553
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1554
fun mk_defname fixes = fixes |> map (fst o fst) |> space_implode "_"
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1555
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1556
fun add_simps fnames post sort extra_qualify label mod_binding moreatts
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1557
  simps lthy =
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1558
  let
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1559
    val spec = post simps
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1560
      |> map (apfst (apsnd (fn ats => moreatts @ ats)))
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1561
      |> map (apfst (apfst extra_qualify))
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1562
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1563
    val (saved_spec_simps, lthy) =
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1564
      fold_map Local_Theory.note spec lthy
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1565
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1566
    val saved_simps = maps snd saved_spec_simps
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1567
    val simps_by_f = sort saved_simps
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1568
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1569
    fun add_for_f fname simps =
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1570
      Local_Theory.note
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1571
        ((mod_binding (Binding.qualify true fname (Binding.name label)), []), simps)
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1572
      #> snd
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1573
  in
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1574
    (saved_simps, fold2 add_for_f fnames simps_by_f lthy)
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1575
  end
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1576
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1577
fun prepare_function is_external prep default_constraint fixspec eqns config lthy =
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1578
  let
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1579
    val constrn_fxs = map (fn (b, T, mx) => (b, SOME (the_default default_constraint T), mx))
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1580
    val ((fixes0, spec0), ctxt') = prep (constrn_fxs fixspec) eqns lthy
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1581
    val fixes = map (apfst (apfst Binding.name_of)) fixes0;
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1582
    val spec = map (fn (bnd, prop) => (bnd, [prop])) spec0;
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1583
    val (eqs, post, sort_cont, cnames) = get_preproc lthy config ctxt' fixes spec
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1584
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1585
    val defname = mk_defname fixes
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1586
    val FunctionConfig {partials, tailrec, default, ...} = config
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1587
    val _ =
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1588
      if tailrec then Output.legacy_feature
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1589
        "'function (tailrec)' command. Use 'partial_function (tailrec)'."
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1590
      else ()
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1591
    val _ =
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1592
      if is_some default then Output.legacy_feature
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1593
        "'function (default)'. Use 'partial_function'."
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1594
      else ()
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1595
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1596
    val ((goal_state, cont), lthy') =
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1597
      prepare_function_mutual config defname fixes eqs lthy
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1598
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1599
    fun afterqed [[proof]] lthy =
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1600
      let
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1601
        val FunctionResult {fs, R, psimps, trsimps, simple_pinducts,
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1602
          termination, domintros, cases, ...} =
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1603
          cont (Thm.close_derivation proof)
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1604
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1605
        val fnames = map (fst o fst) fixes
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1606
        fun qualify n = Binding.name n
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1607
          |> Binding.qualify true defname
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1608
        val conceal_partial = if partials then I else Binding.conceal
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1609
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1610
        val addsmps = add_simps fnames post sort_cont
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1611
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1612
        val (((psimps', pinducts'), (_, [termination'])), lthy) =
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1613
          lthy
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1614
          |> addsmps (conceal_partial o Binding.qualify false "partial")
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1615
               "psimps" conceal_partial psimp_attribs psimps
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1616
          ||> (case trsimps of NONE => I | SOME thms =>
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1617
                   addsmps I "simps" I simp_attribs thms #> snd
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1618
                   #> Spec_Rules.add Spec_Rules.Equational (fs, thms))
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1619
          ||>> Local_Theory.note ((conceal_partial (qualify "pinduct"),
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1620
                 [Attrib.internal (K (Rule_Cases.case_names cnames)),
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1621
                  Attrib.internal (K (Rule_Cases.consumes 1)),
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1622
                  Attrib.internal (K (Induct.induct_pred ""))]), simple_pinducts)
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1623
          ||>> Local_Theory.note ((Binding.conceal (qualify "termination"), []), [termination])
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1624
          ||> (snd o Local_Theory.note ((qualify "cases",
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1625
                 [Attrib.internal (K (Rule_Cases.case_names cnames))]), [cases]))
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1626
          ||> (case domintros of NONE => I | SOME thms => 
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1627
                   Local_Theory.note ((qualify "domintros", []), thms) #> snd)
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1628
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1629
        val info = { add_simps=addsmps, case_names=cnames, psimps=psimps',
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1630
          pinducts=snd pinducts', simps=NONE, inducts=NONE, termination=termination',
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1631
          fs=fs, R=R, defname=defname, is_partial=true }
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1632
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1633
        val _ =
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1634
          if not is_external then ()
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1635
          else Specification.print_consts lthy (K false) (map fst fixes)
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1636
      in
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1637
        (info, 
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1638
         lthy |> Local_Theory.declaration false (add_function_data o morph_function_data info))
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1639
      end
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1640
  in
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1641
    ((goal_state, afterqed), lthy')
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1642
  end
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1643
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1644
*}
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1645
ML {*
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1646
  Proof.theorem;
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1647
  Proof.refine
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1648
*}
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1649
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1650
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1651
ML {*
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1652
fun gen_function is_external prep default_constraint fixspec eqns config lthy =
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1653
  let
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1654
    val ((goal_state, afterqed), lthy') =
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1655
      prepare_function is_external prep default_constraint fixspec eqns config lthy
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1656
  in
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1657
    lthy'
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1658
    |> Proof.theorem NONE (snd oo afterqed) [[(Logic.unprotect (concl_of goal_state), [])]]
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1659
    |> Proof.refine (Method.primitive_text (K goal_state)) 
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1660
    |> Seq.hd
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1661
  end
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1662
*}
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1663
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1664
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1665
ML {*
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1666
val function = gen_function false Specification.check_spec (Type_Infer.anyT HOLogic.typeS)
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1667
val function_cmd = gen_function true Specification.read_spec "_::type"
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1668
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1669
fun add_case_cong n thy =
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1670
  let
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1671
    val cong = #case_cong (Datatype.the_info thy n)
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1672
      |> safe_mk_meta_eq
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1673
  in
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1674
    Context.theory_map
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1675
      (Function_Ctx_Tree.map_function_congs (Thm.add_thm cong)) thy
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1676
  end
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1677
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1678
val setup_case_cong = Datatype.interpretation (K (fold add_case_cong))
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1679
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1680
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1681
(* setup *)
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1682
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1683
val setup =
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1684
  Attrib.setup @{binding fundef_cong}
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1685
    (Attrib.add_del Function_Ctx_Tree.cong_add Function_Ctx_Tree.cong_del)
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1686
    "declaration of congruence rule for function definitions"
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1687
  #> setup_case_cong
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1688
  #> Function_Relation.setup
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1689
  #> Function_Common.Termination_Simps.setup
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1690
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1691
val get_congs = Function_Ctx_Tree.get_function_congs
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1692
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1693
fun get_info ctxt t = Item_Net.retrieve (get_function ctxt) t
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1694
  |> the_single |> snd
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1695
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1696
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1697
(* outer syntax *)
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1698
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1699
val _ =
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1700
  Outer_Syntax.local_theory_to_proof "nominal_primrec" "define recursive functions for nominal types"
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1701
  Keyword.thy_goal
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1702
  (function_parser default_config
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1703
     >> (fn ((config, fixes), statements) => function_cmd fixes statements config))
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1704
*}
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1705
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1706
ML {* trace := true *}
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1707
2654
0f0335d91456 solved subgoals for depth and subst function
Christian Urban <urbanc@in.tum.de>
parents: 2653
diff changeset
  1708
lemma test:
0f0335d91456 solved subgoals for depth and subst function
Christian Urban <urbanc@in.tum.de>
parents: 2653
diff changeset
  1709
  assumes a: "[[x]]lst. t = [[x]]lst. t'"
0f0335d91456 solved subgoals for depth and subst function
Christian Urban <urbanc@in.tum.de>
parents: 2653
diff changeset
  1710
  shows "t = t'"
0f0335d91456 solved subgoals for depth and subst function
Christian Urban <urbanc@in.tum.de>
parents: 2653
diff changeset
  1711
using a
0f0335d91456 solved subgoals for depth and subst function
Christian Urban <urbanc@in.tum.de>
parents: 2653
diff changeset
  1712
apply(simp add: Abs_eq_iff)
0f0335d91456 solved subgoals for depth and subst function
Christian Urban <urbanc@in.tum.de>
parents: 2653
diff changeset
  1713
apply(erule exE)
0f0335d91456 solved subgoals for depth and subst function
Christian Urban <urbanc@in.tum.de>
parents: 2653
diff changeset
  1714
apply(simp only: alphas)
0f0335d91456 solved subgoals for depth and subst function
Christian Urban <urbanc@in.tum.de>
parents: 2653
diff changeset
  1715
apply(erule conjE)+
0f0335d91456 solved subgoals for depth and subst function
Christian Urban <urbanc@in.tum.de>
parents: 2653
diff changeset
  1716
apply(rule sym)
0f0335d91456 solved subgoals for depth and subst function
Christian Urban <urbanc@in.tum.de>
parents: 2653
diff changeset
  1717
apply(clarify)
0f0335d91456 solved subgoals for depth and subst function
Christian Urban <urbanc@in.tum.de>
parents: 2653
diff changeset
  1718
apply(rule supp_perm_eq)
0f0335d91456 solved subgoals for depth and subst function
Christian Urban <urbanc@in.tum.de>
parents: 2653
diff changeset
  1719
apply(subgoal_tac "set [x] \<sharp>* p")
0f0335d91456 solved subgoals for depth and subst function
Christian Urban <urbanc@in.tum.de>
parents: 2653
diff changeset
  1720
apply(auto simp add: fresh_star_def)[1]
0f0335d91456 solved subgoals for depth and subst function
Christian Urban <urbanc@in.tum.de>
parents: 2653
diff changeset
  1721
apply(simp)
0f0335d91456 solved subgoals for depth and subst function
Christian Urban <urbanc@in.tum.de>
parents: 2653
diff changeset
  1722
apply(simp add: fresh_star_def)
0f0335d91456 solved subgoals for depth and subst function
Christian Urban <urbanc@in.tum.de>
parents: 2653
diff changeset
  1723
apply(simp add: fresh_perm)
0f0335d91456 solved subgoals for depth and subst function
Christian Urban <urbanc@in.tum.de>
parents: 2653
diff changeset
  1724
done
0f0335d91456 solved subgoals for depth and subst function
Christian Urban <urbanc@in.tum.de>
parents: 2653
diff changeset
  1725
0f0335d91456 solved subgoals for depth and subst function
Christian Urban <urbanc@in.tum.de>
parents: 2653
diff changeset
  1726
lemma test2:
0f0335d91456 solved subgoals for depth and subst function
Christian Urban <urbanc@in.tum.de>
parents: 2653
diff changeset
  1727
  assumes "a \<sharp> x" "c \<sharp> x" "b \<sharp> y" "c \<sharp> y"
0f0335d91456 solved subgoals for depth and subst function
Christian Urban <urbanc@in.tum.de>
parents: 2653
diff changeset
  1728
  and "(a \<rightleftharpoons> c) \<bullet> x = (b \<rightleftharpoons> c) \<bullet> y"
0f0335d91456 solved subgoals for depth and subst function
Christian Urban <urbanc@in.tum.de>
parents: 2653
diff changeset
  1729
  shows "x = y"
0f0335d91456 solved subgoals for depth and subst function
Christian Urban <urbanc@in.tum.de>
parents: 2653
diff changeset
  1730
using assms by (simp add: swap_fresh_fresh)
0f0335d91456 solved subgoals for depth and subst function
Christian Urban <urbanc@in.tum.de>
parents: 2653
diff changeset
  1731
0f0335d91456 solved subgoals for depth and subst function
Christian Urban <urbanc@in.tum.de>
parents: 2653
diff changeset
  1732
lemma test3:
0f0335d91456 solved subgoals for depth and subst function
Christian Urban <urbanc@in.tum.de>
parents: 2653
diff changeset
  1733
  assumes "x = y"
0f0335d91456 solved subgoals for depth and subst function
Christian Urban <urbanc@in.tum.de>
parents: 2653
diff changeset
  1734
  and "a \<sharp> x" "c \<sharp> x" "b \<sharp> y" "c \<sharp> y"
0f0335d91456 solved subgoals for depth and subst function
Christian Urban <urbanc@in.tum.de>
parents: 2653
diff changeset
  1735
  shows "(a \<rightleftharpoons> c) \<bullet> x = (b \<rightleftharpoons> c) \<bullet> y"
0f0335d91456 solved subgoals for depth and subst function
Christian Urban <urbanc@in.tum.de>
parents: 2653
diff changeset
  1736
using assms by (simp add: swap_fresh_fresh)
0f0335d91456 solved subgoals for depth and subst function
Christian Urban <urbanc@in.tum.de>
parents: 2653
diff changeset
  1737
2649
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1738
nominal_primrec
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1739
  depth :: "lam \<Rightarrow> nat"
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1740
where
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1741
  "depth (Var x) = 1"
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1742
| "depth (App t1 t2) = (max (depth t1) (depth t2)) + 1"
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1743
| "depth (Lam x t) = (depth t) + 1"
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1744
apply(rule_tac y="x" in lam.exhaust)
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1745
apply(simp_all)[3]
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1746
apply(simp_all only: lam.distinct)
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1747
apply(simp add: lam.eq_iff)
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1748
apply(simp add: lam.eq_iff)
2660
3342a2d13d95 nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
parents: 2657
diff changeset
  1749
apply(subst (asm) Abs_eq_iff)
3342a2d13d95 nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
parents: 2657
diff changeset
  1750
apply(erule exE)
3342a2d13d95 nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
parents: 2657
diff changeset
  1751
apply(simp add: alphas)
3342a2d13d95 nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
parents: 2657
diff changeset
  1752
apply(clarify)
3342a2d13d95 nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
parents: 2657
diff changeset
  1753
oops
3342a2d13d95 nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
parents: 2657
diff changeset
  1754
3342a2d13d95 nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
parents: 2657
diff changeset
  1755
lemma removeAll_eqvt[eqvt]:
3342a2d13d95 nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
parents: 2657
diff changeset
  1756
  shows "p \<bullet> (removeAll x xs) = removeAll (p \<bullet> x) (p \<bullet> xs)"
3342a2d13d95 nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
parents: 2657
diff changeset
  1757
by (induct xs) (auto)
3342a2d13d95 nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
parents: 2657
diff changeset
  1758
3342a2d13d95 nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
parents: 2657
diff changeset
  1759
nominal_primrec 
3342a2d13d95 nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
parents: 2657
diff changeset
  1760
  frees_lst :: "lam \<Rightarrow> atom list"
3342a2d13d95 nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
parents: 2657
diff changeset
  1761
where
3342a2d13d95 nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
parents: 2657
diff changeset
  1762
  "frees_lst (Var x) = [atom x]"
3342a2d13d95 nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
parents: 2657
diff changeset
  1763
| "frees_lst (App t1 t2) = (frees_lst t1) @ (frees_lst t2)"
3342a2d13d95 nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
parents: 2657
diff changeset
  1764
| "frees_lst (Lam x t) = removeAll (atom x) (frees_lst t)"
3342a2d13d95 nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
parents: 2657
diff changeset
  1765
apply(rule_tac y="x" in lam.exhaust)
3342a2d13d95 nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
parents: 2657
diff changeset
  1766
apply(simp_all)[3]
3342a2d13d95 nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
parents: 2657
diff changeset
  1767
apply(simp_all only: lam.distinct)
3342a2d13d95 nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
parents: 2657
diff changeset
  1768
apply(simp add: lam.eq_iff)
3342a2d13d95 nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
parents: 2657
diff changeset
  1769
apply(simp add: lam.eq_iff)
3342a2d13d95 nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
parents: 2657
diff changeset
  1770
apply(simp add: lam.eq_iff)
3342a2d13d95 nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
parents: 2657
diff changeset
  1771
apply(simp add: Abs_eq_iff)
3342a2d13d95 nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
parents: 2657
diff changeset
  1772
apply(erule exE)
3342a2d13d95 nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
parents: 2657
diff changeset
  1773
apply(simp add: alphas)
3342a2d13d95 nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
parents: 2657
diff changeset
  1774
apply(simp add: atom_eqvt)
3342a2d13d95 nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
parents: 2657
diff changeset
  1775
apply(clarify)
3342a2d13d95 nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
parents: 2657
diff changeset
  1776
oops
3342a2d13d95 nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
parents: 2657
diff changeset
  1777
3342a2d13d95 nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
parents: 2657
diff changeset
  1778
nominal_primrec
3342a2d13d95 nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
parents: 2657
diff changeset
  1779
  subst :: "lam \<Rightarrow> name \<Rightarrow> lam \<Rightarrow> lam"  ("_ [_ ::= _]" [100,100,100] 100)
3342a2d13d95 nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
parents: 2657
diff changeset
  1780
where
3342a2d13d95 nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
parents: 2657
diff changeset
  1781
  "(Var x)[y ::= s] = (if x=y then s else (Var x))"
3342a2d13d95 nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
parents: 2657
diff changeset
  1782
| "(App t\<^isub>1 t\<^isub>2)[y ::= s] = App (t\<^isub>1[y ::= s]) (t\<^isub>2[y ::= s])"
3342a2d13d95 nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
parents: 2657
diff changeset
  1783
| "atom x \<sharp> (y, s) \<Longrightarrow> (Lam x t)[y ::= s] = Lam x (t[y ::= s])"
3342a2d13d95 nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
parents: 2657
diff changeset
  1784
apply(case_tac x)
3342a2d13d95 nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
parents: 2657
diff changeset
  1785
apply(simp)
3342a2d13d95 nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
parents: 2657
diff changeset
  1786
apply(rule_tac y="a" and c="(b, c)" in lam.strong_exhaust)
3342a2d13d95 nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
parents: 2657
diff changeset
  1787
apply(simp add: lam.eq_iff lam.distinct)
3342a2d13d95 nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
parents: 2657
diff changeset
  1788
apply(auto)[1]
3342a2d13d95 nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
parents: 2657
diff changeset
  1789
apply(simp add: lam.eq_iff lam.distinct)
3342a2d13d95 nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
parents: 2657
diff changeset
  1790
apply(auto)[1]
3342a2d13d95 nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
parents: 2657
diff changeset
  1791
apply(simp add: fresh_star_def lam.eq_iff lam.distinct)
3342a2d13d95 nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
parents: 2657
diff changeset
  1792
apply(simp_all add: lam.distinct)[5]
3342a2d13d95 nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
parents: 2657
diff changeset
  1793
apply(simp add: lam.eq_iff)
3342a2d13d95 nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
parents: 2657
diff changeset
  1794
apply(simp add: lam.eq_iff)
3342a2d13d95 nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
parents: 2657
diff changeset
  1795
apply(simp add: lam.eq_iff)
3342a2d13d95 nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
parents: 2657
diff changeset
  1796
apply(erule conjE)+
2661
16896fd8eff5 subst also works now
Christian Urban <urbanc@in.tum.de>
parents: 2660
diff changeset
  1797
oops
2660
3342a2d13d95 nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
parents: 2657
diff changeset
  1798
3342a2d13d95 nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
parents: 2657
diff changeset
  1799
3342a2d13d95 nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
parents: 2657
diff changeset
  1800
3342a2d13d95 nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
parents: 2657
diff changeset
  1801
nominal_primrec
3342a2d13d95 nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
parents: 2657
diff changeset
  1802
  depth :: "lam \<Rightarrow> nat"
3342a2d13d95 nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
parents: 2657
diff changeset
  1803
where
3342a2d13d95 nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
parents: 2657
diff changeset
  1804
  "depth (Var x) = 1"
3342a2d13d95 nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
parents: 2657
diff changeset
  1805
| "depth (App t1 t2) = (max (depth t1) (depth t2)) + 1"
3342a2d13d95 nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
parents: 2657
diff changeset
  1806
| "depth (Lam x t) = (depth t) + 1"
3342a2d13d95 nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
parents: 2657
diff changeset
  1807
apply(rule_tac y="x" in lam.exhaust)
3342a2d13d95 nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
parents: 2657
diff changeset
  1808
apply(simp_all)[3]
3342a2d13d95 nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
parents: 2657
diff changeset
  1809
apply(simp_all only: lam.distinct)
3342a2d13d95 nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
parents: 2657
diff changeset
  1810
apply(simp add: lam.eq_iff)
3342a2d13d95 nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
parents: 2657
diff changeset
  1811
apply(simp add: lam.eq_iff)
3342a2d13d95 nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
parents: 2657
diff changeset
  1812
(*
3342a2d13d95 nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
parents: 2657
diff changeset
  1813
apply(subst (asm) Abs_eq_iff)
3342a2d13d95 nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
parents: 2657
diff changeset
  1814
apply(erule exE)
3342a2d13d95 nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
parents: 2657
diff changeset
  1815
apply(simp add: alphas)
3342a2d13d95 nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
parents: 2657
diff changeset
  1816
apply(clarify)
3342a2d13d95 nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
parents: 2657
diff changeset
  1817
*)
2654
0f0335d91456 solved subgoals for depth and subst function
Christian Urban <urbanc@in.tum.de>
parents: 2653
diff changeset
  1818
apply(subgoal_tac "finite (supp (x, xa, t, ta, depth_sumC t, depth_sumC ta))")
0f0335d91456 solved subgoals for depth and subst function
Christian Urban <urbanc@in.tum.de>
parents: 2653
diff changeset
  1819
apply(erule_tac ?'a="name" in obtain_at_base)
0f0335d91456 solved subgoals for depth and subst function
Christian Urban <urbanc@in.tum.de>
parents: 2653
diff changeset
  1820
unfolding fresh_def[symmetric]
0f0335d91456 solved subgoals for depth and subst function
Christian Urban <urbanc@in.tum.de>
parents: 2653
diff changeset
  1821
apply(drule_tac a="atom x" and b="atom xa" and c="atom a" in test3)
0f0335d91456 solved subgoals for depth and subst function
Christian Urban <urbanc@in.tum.de>
parents: 2653
diff changeset
  1822
apply(simp add: Abs_fresh_iff)
0f0335d91456 solved subgoals for depth and subst function
Christian Urban <urbanc@in.tum.de>
parents: 2653
diff changeset
  1823
apply(simp add: Abs_fresh_iff)
0f0335d91456 solved subgoals for depth and subst function
Christian Urban <urbanc@in.tum.de>
parents: 2653
diff changeset
  1824
apply(simp add: Abs_fresh_iff)
0f0335d91456 solved subgoals for depth and subst function
Christian Urban <urbanc@in.tum.de>
parents: 2653
diff changeset
  1825
apply(simp add: Abs_fresh_iff)
0f0335d91456 solved subgoals for depth and subst function
Christian Urban <urbanc@in.tum.de>
parents: 2653
diff changeset
  1826
apply(rule_tac a="atom x" and b="atom xa" and c="atom a" in test2)
0f0335d91456 solved subgoals for depth and subst function
Christian Urban <urbanc@in.tum.de>
parents: 2653
diff changeset
  1827
apply(simp add: pure_fresh)
0f0335d91456 solved subgoals for depth and subst function
Christian Urban <urbanc@in.tum.de>
parents: 2653
diff changeset
  1828
apply(simp add: pure_fresh)
0f0335d91456 solved subgoals for depth and subst function
Christian Urban <urbanc@in.tum.de>
parents: 2653
diff changeset
  1829
apply(simp add: pure_fresh)
0f0335d91456 solved subgoals for depth and subst function
Christian Urban <urbanc@in.tum.de>
parents: 2653
diff changeset
  1830
apply(simp add: pure_fresh)
0f0335d91456 solved subgoals for depth and subst function
Christian Urban <urbanc@in.tum.de>
parents: 2653
diff changeset
  1831
apply(simp add: eqvt_at_def)
0f0335d91456 solved subgoals for depth and subst function
Christian Urban <urbanc@in.tum.de>
parents: 2653
diff changeset
  1832
apply(drule test)
0f0335d91456 solved subgoals for depth and subst function
Christian Urban <urbanc@in.tum.de>
parents: 2653
diff changeset
  1833
apply(simp)
0f0335d91456 solved subgoals for depth and subst function
Christian Urban <urbanc@in.tum.de>
parents: 2653
diff changeset
  1834
apply(simp add: finite_supp)
0f0335d91456 solved subgoals for depth and subst function
Christian Urban <urbanc@in.tum.de>
parents: 2653
diff changeset
  1835
done
0f0335d91456 solved subgoals for depth and subst function
Christian Urban <urbanc@in.tum.de>
parents: 2653
diff changeset
  1836
0f0335d91456 solved subgoals for depth and subst function
Christian Urban <urbanc@in.tum.de>
parents: 2653
diff changeset
  1837
termination depth
0f0335d91456 solved subgoals for depth and subst function
Christian Urban <urbanc@in.tum.de>
parents: 2653
diff changeset
  1838
  apply(relation "measure size")
0f0335d91456 solved subgoals for depth and subst function
Christian Urban <urbanc@in.tum.de>
parents: 2653
diff changeset
  1839
  apply(auto simp add: lam.size)
0f0335d91456 solved subgoals for depth and subst function
Christian Urban <urbanc@in.tum.de>
parents: 2653
diff changeset
  1840
  done
2649
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1841
2651
4aa72a88b2c1 equivariance of THE_default under the uniqueness assumption
Christian Urban <urbanc@in.tum.de>
parents: 2650
diff changeset
  1842
thm depth.psimps
2654
0f0335d91456 solved subgoals for depth and subst function
Christian Urban <urbanc@in.tum.de>
parents: 2653
diff changeset
  1843
thm depth.simps
2651
4aa72a88b2c1 equivariance of THE_default under the uniqueness assumption
Christian Urban <urbanc@in.tum.de>
parents: 2650
diff changeset
  1844
2657
1ea9c059fc0f a few lemmas about freshness for at and at_base
Christian Urban <urbanc@in.tum.de>
parents: 2656
diff changeset
  1845
1ea9c059fc0f a few lemmas about freshness for at and at_base
Christian Urban <urbanc@in.tum.de>
parents: 2656
diff changeset
  1846
lemma swap_set_not_in_at:
1ea9c059fc0f a few lemmas about freshness for at and at_base
Christian Urban <urbanc@in.tum.de>
parents: 2656
diff changeset
  1847
  fixes a b::"'a::at"
1ea9c059fc0f a few lemmas about freshness for at and at_base
Christian Urban <urbanc@in.tum.de>
parents: 2656
diff changeset
  1848
  and   S::"'a::at set"
1ea9c059fc0f a few lemmas about freshness for at and at_base
Christian Urban <urbanc@in.tum.de>
parents: 2656
diff changeset
  1849
  assumes a: "a \<notin> S" "b \<notin> S"
1ea9c059fc0f a few lemmas about freshness for at and at_base
Christian Urban <urbanc@in.tum.de>
parents: 2656
diff changeset
  1850
  shows "(a \<leftrightarrow> b) \<bullet> S = S"
1ea9c059fc0f a few lemmas about freshness for at and at_base
Christian Urban <urbanc@in.tum.de>
parents: 2656
diff changeset
  1851
  unfolding permute_set_eq
1ea9c059fc0f a few lemmas about freshness for at and at_base
Christian Urban <urbanc@in.tum.de>
parents: 2656
diff changeset
  1852
  using a by (auto simp add: permute_flip_at)
1ea9c059fc0f a few lemmas about freshness for at and at_base
Christian Urban <urbanc@in.tum.de>
parents: 2656
diff changeset
  1853
2660
3342a2d13d95 nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
parents: 2657
diff changeset
  1854
lemma removeAll_eqvt[eqvt]:
3342a2d13d95 nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
parents: 2657
diff changeset
  1855
  shows "p \<bullet> (removeAll x xs) = removeAll (p \<bullet> x) (p \<bullet> xs)"
3342a2d13d95 nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
parents: 2657
diff changeset
  1856
by (induct xs) (auto)
3342a2d13d95 nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
parents: 2657
diff changeset
  1857
3342a2d13d95 nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
parents: 2657
diff changeset
  1858
nominal_primrec 
3342a2d13d95 nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
parents: 2657
diff changeset
  1859
  frees_lst :: "lam \<Rightarrow> atom list"
3342a2d13d95 nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
parents: 2657
diff changeset
  1860
where
3342a2d13d95 nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
parents: 2657
diff changeset
  1861
  "frees_lst (Var x) = [atom x]"
3342a2d13d95 nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
parents: 2657
diff changeset
  1862
| "frees_lst (App t1 t2) = (frees_lst t1) @ (frees_lst t2)"
3342a2d13d95 nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
parents: 2657
diff changeset
  1863
| "frees_lst (Lam x t) = removeAll (atom x) (frees_lst t)"
3342a2d13d95 nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
parents: 2657
diff changeset
  1864
apply(rule_tac y="x" in lam.exhaust)
3342a2d13d95 nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
parents: 2657
diff changeset
  1865
apply(simp_all)[3]
3342a2d13d95 nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
parents: 2657
diff changeset
  1866
apply(simp_all only: lam.distinct)
3342a2d13d95 nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
parents: 2657
diff changeset
  1867
apply(simp add: lam.eq_iff)
3342a2d13d95 nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
parents: 2657
diff changeset
  1868
apply(simp add: lam.eq_iff)
3342a2d13d95 nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
parents: 2657
diff changeset
  1869
apply(simp add: lam.eq_iff)
3342a2d13d95 nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
parents: 2657
diff changeset
  1870
apply(simp add: Abs_eq_iff)
3342a2d13d95 nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
parents: 2657
diff changeset
  1871
apply(erule exE)
3342a2d13d95 nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
parents: 2657
diff changeset
  1872
apply(simp add: alphas)
3342a2d13d95 nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
parents: 2657
diff changeset
  1873
apply(simp add: atom_eqvt)
3342a2d13d95 nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
parents: 2657
diff changeset
  1874
apply(clarify)
3342a2d13d95 nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
parents: 2657
diff changeset
  1875
apply(rule trans)
3342a2d13d95 nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
parents: 2657
diff changeset
  1876
apply(rule sym)
3342a2d13d95 nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
parents: 2657
diff changeset
  1877
apply(rule_tac p="p" in supp_perm_eq)
3342a2d13d95 nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
parents: 2657
diff changeset
  1878
oops
2657
1ea9c059fc0f a few lemmas about freshness for at and at_base
Christian Urban <urbanc@in.tum.de>
parents: 2656
diff changeset
  1879
2649
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1880
nominal_primrec 
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1881
  frees :: "lam \<Rightarrow> name set"
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1882
where
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1883
  "frees (Var x) = {x}"
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1884
| "frees (App t1 t2) = (frees t1) \<union> (frees t2)"
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1885
| "frees (Lam x t) = (frees t) - {x}"
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1886
apply(rule_tac y="x" in lam.exhaust)
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1887
apply(simp_all)[3]
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1888
apply(simp_all only: lam.distinct)
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1889
apply(simp add: lam.eq_iff)
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1890
apply(simp add: lam.eq_iff)
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1891
apply(simp add: lam.eq_iff)
2654
0f0335d91456 solved subgoals for depth and subst function
Christian Urban <urbanc@in.tum.de>
parents: 2653
diff changeset
  1892
apply(subgoal_tac "finite (supp (x, xa, t, ta, frees_sumC t, frees_sumC ta))")
0f0335d91456 solved subgoals for depth and subst function
Christian Urban <urbanc@in.tum.de>
parents: 2653
diff changeset
  1893
apply(erule_tac ?'a="name" in obtain_at_base)
0f0335d91456 solved subgoals for depth and subst function
Christian Urban <urbanc@in.tum.de>
parents: 2653
diff changeset
  1894
unfolding fresh_def[symmetric]
0f0335d91456 solved subgoals for depth and subst function
Christian Urban <urbanc@in.tum.de>
parents: 2653
diff changeset
  1895
apply(drule_tac a="atom x" and b="atom xa" and c="atom a" in test3)
0f0335d91456 solved subgoals for depth and subst function
Christian Urban <urbanc@in.tum.de>
parents: 2653
diff changeset
  1896
apply(simp add: Abs_fresh_iff)
0f0335d91456 solved subgoals for depth and subst function
Christian Urban <urbanc@in.tum.de>
parents: 2653
diff changeset
  1897
apply(simp add: Abs_fresh_iff)
0f0335d91456 solved subgoals for depth and subst function
Christian Urban <urbanc@in.tum.de>
parents: 2653
diff changeset
  1898
apply(simp add: Abs_fresh_iff)
0f0335d91456 solved subgoals for depth and subst function
Christian Urban <urbanc@in.tum.de>
parents: 2653
diff changeset
  1899
apply(simp add: Abs_fresh_iff)
0f0335d91456 solved subgoals for depth and subst function
Christian Urban <urbanc@in.tum.de>
parents: 2653
diff changeset
  1900
apply(simp)
0f0335d91456 solved subgoals for depth and subst function
Christian Urban <urbanc@in.tum.de>
parents: 2653
diff changeset
  1901
apply(drule test)
2657
1ea9c059fc0f a few lemmas about freshness for at and at_base
Christian Urban <urbanc@in.tum.de>
parents: 2656
diff changeset
  1902
apply(rule_tac t="frees_sumC t - {x}" and s="(x \<leftrightarrow> a) \<bullet> (frees_sumC t - {x})" in subst)
2654
0f0335d91456 solved subgoals for depth and subst function
Christian Urban <urbanc@in.tum.de>
parents: 2653
diff changeset
  1903
oops
2649
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1904
2660
3342a2d13d95 nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
parents: 2657
diff changeset
  1905
thm Abs_eq_iff[simplified alphas]
3342a2d13d95 nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
parents: 2657
diff changeset
  1906
3342a2d13d95 nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
parents: 2657
diff changeset
  1907
lemma Abs_set_eq_iff2:
3342a2d13d95 nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
parents: 2657
diff changeset
  1908
  fixes x y::"'a::fs"
3342a2d13d95 nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
parents: 2657
diff changeset
  1909
  shows "[bs]set. x = [cs]set. y \<longleftrightarrow>
3342a2d13d95 nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
parents: 2657
diff changeset
  1910
    (\<exists>p. supp ([bs]set. x) = supp ([cs]set. y) \<and>
3342a2d13d95 nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
parents: 2657
diff changeset
  1911
         supp ([bs]set. x) \<sharp>* p \<and>
3342a2d13d95 nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
parents: 2657
diff changeset
  1912
         p \<bullet> x = y \<and> p \<bullet> bs = cs)"
3342a2d13d95 nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
parents: 2657
diff changeset
  1913
unfolding Abs_eq_iff
3342a2d13d95 nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
parents: 2657
diff changeset
  1914
unfolding alphas
3342a2d13d95 nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
parents: 2657
diff changeset
  1915
unfolding supp_Abs
3342a2d13d95 nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
parents: 2657
diff changeset
  1916
by simp
3342a2d13d95 nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
parents: 2657
diff changeset
  1917
3342a2d13d95 nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
parents: 2657
diff changeset
  1918
lemma Abs_set_eq_iff3:
3342a2d13d95 nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
parents: 2657
diff changeset
  1919
  fixes x y::"'a::fs"
3342a2d13d95 nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
parents: 2657
diff changeset
  1920
  assumes a: "finite bs" "finite cs"
3342a2d13d95 nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
parents: 2657
diff changeset
  1921
  shows "[bs]set. x = [cs]set. y \<longleftrightarrow>
3342a2d13d95 nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
parents: 2657
diff changeset
  1922
    (\<exists>p. supp ([bs]set. x) = supp ([cs]set. y) \<and>
3342a2d13d95 nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
parents: 2657
diff changeset
  1923
         supp ([bs]set. x) \<sharp>* p \<and>
3342a2d13d95 nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
parents: 2657
diff changeset
  1924
         p \<bullet> x = y \<and> p \<bullet> bs = cs \<and>
3342a2d13d95 nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
parents: 2657
diff changeset
  1925
         supp p \<subseteq> bs \<union> cs)"
3342a2d13d95 nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
parents: 2657
diff changeset
  1926
unfolding Abs_eq_iff
3342a2d13d95 nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
parents: 2657
diff changeset
  1927
unfolding alphas
3342a2d13d95 nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
parents: 2657
diff changeset
  1928
unfolding supp_Abs
3342a2d13d95 nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
parents: 2657
diff changeset
  1929
apply(auto)
3342a2d13d95 nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
parents: 2657
diff changeset
  1930
using set_renaming_perm
3342a2d13d95 nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
parents: 2657
diff changeset
  1931
sorry
3342a2d13d95 nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
parents: 2657
diff changeset
  1932
3342a2d13d95 nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
parents: 2657
diff changeset
  1933
lemma Abs_eq_iff2:
3342a2d13d95 nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
parents: 2657
diff changeset
  1934
  fixes x y::"'a::fs"
3342a2d13d95 nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
parents: 2657
diff changeset
  1935
  shows "[bs]lst. x = [cs]lst. y \<longleftrightarrow>
3342a2d13d95 nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
parents: 2657
diff changeset
  1936
    (\<exists>p. supp ([bs]lst. x) = supp ([cs]lst. y) \<and>
3342a2d13d95 nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
parents: 2657
diff changeset
  1937
         supp ([bs]lst. x) \<sharp>* p \<and>
3342a2d13d95 nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
parents: 2657
diff changeset
  1938
         p \<bullet> x = y \<and> p \<bullet> bs = cs)"
3342a2d13d95 nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
parents: 2657
diff changeset
  1939
unfolding Abs_eq_iff
3342a2d13d95 nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
parents: 2657
diff changeset
  1940
unfolding alphas
3342a2d13d95 nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
parents: 2657
diff changeset
  1941
unfolding supp_Abs
3342a2d13d95 nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
parents: 2657
diff changeset
  1942
by simp
3342a2d13d95 nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
parents: 2657
diff changeset
  1943
3342a2d13d95 nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
parents: 2657
diff changeset
  1944
lemma Abs_eq_iff3:
3342a2d13d95 nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
parents: 2657
diff changeset
  1945
  fixes x y::"'a::fs"
3342a2d13d95 nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
parents: 2657
diff changeset
  1946
  shows "[bs]lst. x = [cs]lst. y \<longleftrightarrow>
3342a2d13d95 nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
parents: 2657
diff changeset
  1947
    (\<exists>p. supp ([bs]lst. x) = supp ([cs]lst. y) \<and>
3342a2d13d95 nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
parents: 2657
diff changeset
  1948
         supp ([bs]lst. x) \<sharp>* p \<and>
3342a2d13d95 nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
parents: 2657
diff changeset
  1949
         p \<bullet> x = y \<and> p \<bullet> bs = cs \<and>
3342a2d13d95 nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
parents: 2657
diff changeset
  1950
         supp p \<subseteq> set bs \<union> set cs)"
3342a2d13d95 nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
parents: 2657
diff changeset
  1951
unfolding Abs_eq_iff
3342a2d13d95 nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
parents: 2657
diff changeset
  1952
unfolding alphas
3342a2d13d95 nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
parents: 2657
diff changeset
  1953
unfolding supp_Abs
3342a2d13d95 nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
parents: 2657
diff changeset
  1954
apply(auto)
3342a2d13d95 nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
parents: 2657
diff changeset
  1955
using list_renaming_perm
3342a2d13d95 nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
parents: 2657
diff changeset
  1956
apply -
3342a2d13d95 nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
parents: 2657
diff changeset
  1957
apply(drule_tac x="bs" in meta_spec)
3342a2d13d95 nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
parents: 2657
diff changeset
  1958
apply(drule_tac x="p" in meta_spec)
3342a2d13d95 nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
parents: 2657
diff changeset
  1959
apply(erule exE)
3342a2d13d95 nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
parents: 2657
diff changeset
  1960
apply(rule_tac x="q" in exI)
3342a2d13d95 nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
parents: 2657
diff changeset
  1961
apply(simp add: set_eqvt)
3342a2d13d95 nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
parents: 2657
diff changeset
  1962
sorry
3342a2d13d95 nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
parents: 2657
diff changeset
  1963
3342a2d13d95 nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
parents: 2657
diff changeset
  1964
nominal_primrec
3342a2d13d95 nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
parents: 2657
diff changeset
  1965
  subst :: "lam \<Rightarrow> name \<Rightarrow> lam \<Rightarrow> lam"  ("_ [_ ::= _]" [100,100,100] 100)
3342a2d13d95 nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
parents: 2657
diff changeset
  1966
where
3342a2d13d95 nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
parents: 2657
diff changeset
  1967
  "(Var x)[y ::= s] = (if x=y then s else (Var x))"
3342a2d13d95 nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
parents: 2657
diff changeset
  1968
| "(App t\<^isub>1 t\<^isub>2)[y ::= s] = App (t\<^isub>1[y ::= s]) (t\<^isub>2[y ::= s])"
3342a2d13d95 nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
parents: 2657
diff changeset
  1969
| "atom x \<sharp> (y, s) \<Longrightarrow> (Lam x t)[y ::= s] = Lam x (t[y ::= s])"
3342a2d13d95 nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
parents: 2657
diff changeset
  1970
apply(case_tac x)
3342a2d13d95 nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
parents: 2657
diff changeset
  1971
apply(simp)
3342a2d13d95 nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
parents: 2657
diff changeset
  1972
apply(rule_tac y="a" and c="(b, c)" in lam.strong_exhaust)
3342a2d13d95 nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
parents: 2657
diff changeset
  1973
apply(simp add: lam.eq_iff lam.distinct)
3342a2d13d95 nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
parents: 2657
diff changeset
  1974
apply(auto)[1]
3342a2d13d95 nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
parents: 2657
diff changeset
  1975
apply(simp add: lam.eq_iff lam.distinct)
3342a2d13d95 nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
parents: 2657
diff changeset
  1976
apply(auto)[1]
3342a2d13d95 nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
parents: 2657
diff changeset
  1977
apply(simp add: fresh_star_def lam.eq_iff lam.distinct)
3342a2d13d95 nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
parents: 2657
diff changeset
  1978
apply(simp_all add: lam.distinct)[5]
3342a2d13d95 nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
parents: 2657
diff changeset
  1979
apply(simp add: lam.eq_iff)
3342a2d13d95 nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
parents: 2657
diff changeset
  1980
apply(simp add: lam.eq_iff)
3342a2d13d95 nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
parents: 2657
diff changeset
  1981
apply(simp add: lam.eq_iff)
3342a2d13d95 nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
parents: 2657
diff changeset
  1982
apply(erule conjE)+
3342a2d13d95 nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
parents: 2657
diff changeset
  1983
apply(subst (asm) Abs_eq_iff3) 
3342a2d13d95 nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
parents: 2657
diff changeset
  1984
apply(erule exE)
3342a2d13d95 nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
parents: 2657
diff changeset
  1985
apply(erule conjE)+
3342a2d13d95 nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
parents: 2657
diff changeset
  1986
apply(clarify)
3342a2d13d95 nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
parents: 2657
diff changeset
  1987
apply(perm_simp)
3342a2d13d95 nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
parents: 2657
diff changeset
  1988
apply(simp)
3342a2d13d95 nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
parents: 2657
diff changeset
  1989
apply(rule trans)
3342a2d13d95 nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
parents: 2657
diff changeset
  1990
apply(rule sym)
3342a2d13d95 nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
parents: 2657
diff changeset
  1991
apply(rule_tac p="p" in supp_perm_eq)
3342a2d13d95 nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
parents: 2657
diff changeset
  1992
apply(rule fresh_star_supp_conv)
3342a2d13d95 nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
parents: 2657
diff changeset
  1993
apply(drule fresh_star_supp_conv)
3342a2d13d95 nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
parents: 2657
diff changeset
  1994
apply(simp add: Abs_fresh_star_iff)
3342a2d13d95 nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
parents: 2657
diff changeset
  1995
thm fresh_eqvt_at
3342a2d13d95 nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
parents: 2657
diff changeset
  1996
apply(rule_tac f="subst_sumC" in fresh_eqvt_at)
3342a2d13d95 nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
parents: 2657
diff changeset
  1997
apply(assumption)
3342a2d13d95 nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
parents: 2657
diff changeset
  1998
apply(simp add: finite_supp)
3342a2d13d95 nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
parents: 2657
diff changeset
  1999
prefer 2
3342a2d13d95 nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
parents: 2657
diff changeset
  2000
apply(simp)
3342a2d13d95 nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
parents: 2657
diff changeset
  2001
apply(simp add: eqvt_at_def)
3342a2d13d95 nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
parents: 2657
diff changeset
  2002
apply(perm_simp)
3342a2d13d95 nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
parents: 2657
diff changeset
  2003
apply(subgoal_tac "p \<bullet> ya = ya")
3342a2d13d95 nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
parents: 2657
diff changeset
  2004
apply(subgoal_tac "p \<bullet> sa = sa")
3342a2d13d95 nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
parents: 2657
diff changeset
  2005
apply(simp)
3342a2d13d95 nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
parents: 2657
diff changeset
  2006
apply(rule supp_perm_eq)
3342a2d13d95 nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
parents: 2657
diff changeset
  2007
apply(rule fresh_star_supp_conv)
3342a2d13d95 nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
parents: 2657
diff changeset
  2008
apply(auto simp add: fresh_star_def fresh_Pair)[1]
3342a2d13d95 nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
parents: 2657
diff changeset
  2009
apply(rule supp_perm_eq)
3342a2d13d95 nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
parents: 2657
diff changeset
  2010
apply(rule fresh_star_supp_conv)
3342a2d13d95 nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
parents: 2657
diff changeset
  2011
apply(auto simp add: fresh_star_def fresh_Pair)[1]
3342a2d13d95 nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
parents: 2657
diff changeset
  2012
3342a2d13d95 nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
parents: 2657
diff changeset
  2013
3342a2d13d95 nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
parents: 2657
diff changeset
  2014
2649
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  2015
nominal_primrec
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  2016
  subst :: "lam \<Rightarrow> name \<Rightarrow> lam \<Rightarrow> lam"  ("_ [_ ::= _]" [100,100,100] 100)
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  2017
where
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  2018
  "(Var x)[y ::= s] = (if x=y then s else (Var x))"
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  2019
| "(App t\<^isub>1 t\<^isub>2)[y ::= s] = App (t\<^isub>1[y ::= s]) (t\<^isub>2[y ::= s])"
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  2020
| "atom x \<sharp> (y, s) \<Longrightarrow> (Lam x t)[y ::= s] = Lam x (t[y ::= s])"
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  2021
apply(case_tac x)
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  2022
apply(simp)
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  2023
apply(rule_tac y="a" and c="(b, c)" in lam.strong_exhaust)
2653
d0f774d06e6e added eqvt_at premises in function definition - however not proved at the moment
Christian Urban <urbanc@in.tum.de>
parents: 2652
diff changeset
  2024
apply(simp add: lam.eq_iff lam.distinct)
2649
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  2025
apply(auto)[1]
2653
d0f774d06e6e added eqvt_at premises in function definition - however not proved at the moment
Christian Urban <urbanc@in.tum.de>
parents: 2652
diff changeset
  2026
apply(simp add: lam.eq_iff lam.distinct)
d0f774d06e6e added eqvt_at premises in function definition - however not proved at the moment
Christian Urban <urbanc@in.tum.de>
parents: 2652
diff changeset
  2027
apply(auto)[1]
d0f774d06e6e added eqvt_at premises in function definition - however not proved at the moment
Christian Urban <urbanc@in.tum.de>
parents: 2652
diff changeset
  2028
apply(simp add: fresh_star_def lam.eq_iff lam.distinct)
2654
0f0335d91456 solved subgoals for depth and subst function
Christian Urban <urbanc@in.tum.de>
parents: 2653
diff changeset
  2029
apply(simp_all add: lam.distinct)[5]
2649
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  2030
apply(simp add: lam.eq_iff)
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  2031
apply(simp add: lam.eq_iff)
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  2032
apply(simp add: lam.eq_iff)
2654
0f0335d91456 solved subgoals for depth and subst function
Christian Urban <urbanc@in.tum.de>
parents: 2653
diff changeset
  2033
apply(subgoal_tac "finite (supp (ya, sa, x, xa, t, ta, subst_sumC (t, ya, sa), subst_sumC (ta, ya, sa)))")
0f0335d91456 solved subgoals for depth and subst function
Christian Urban <urbanc@in.tum.de>
parents: 2653
diff changeset
  2034
apply(erule_tac ?'a="name" in obtain_at_base)
0f0335d91456 solved subgoals for depth and subst function
Christian Urban <urbanc@in.tum.de>
parents: 2653
diff changeset
  2035
unfolding fresh_def[symmetric]
0f0335d91456 solved subgoals for depth and subst function
Christian Urban <urbanc@in.tum.de>
parents: 2653
diff changeset
  2036
apply(rule_tac a="atom x" and b="atom xa" and c="atom a" in test2)
0f0335d91456 solved subgoals for depth and subst function
Christian Urban <urbanc@in.tum.de>
parents: 2653
diff changeset
  2037
apply(simp add: Abs_fresh_iff)
0f0335d91456 solved subgoals for depth and subst function
Christian Urban <urbanc@in.tum.de>
parents: 2653
diff changeset
  2038
apply(simp add: Abs_fresh_iff)
0f0335d91456 solved subgoals for depth and subst function
Christian Urban <urbanc@in.tum.de>
parents: 2653
diff changeset
  2039
apply(simp add: Abs_fresh_iff)
0f0335d91456 solved subgoals for depth and subst function
Christian Urban <urbanc@in.tum.de>
parents: 2653
diff changeset
  2040
apply(simp add: Abs_fresh_iff)
0f0335d91456 solved subgoals for depth and subst function
Christian Urban <urbanc@in.tum.de>
parents: 2653
diff changeset
  2041
apply(erule conjE)+
0f0335d91456 solved subgoals for depth and subst function
Christian Urban <urbanc@in.tum.de>
parents: 2653
diff changeset
  2042
apply(drule_tac a="atom x" and b="atom xa" and c="atom a" in test3)
0f0335d91456 solved subgoals for depth and subst function
Christian Urban <urbanc@in.tum.de>
parents: 2653
diff changeset
  2043
apply(simp add: Abs_fresh_iff)
0f0335d91456 solved subgoals for depth and subst function
Christian Urban <urbanc@in.tum.de>
parents: 2653
diff changeset
  2044
apply(simp add: Abs_fresh_iff)
0f0335d91456 solved subgoals for depth and subst function
Christian Urban <urbanc@in.tum.de>
parents: 2653
diff changeset
  2045
apply(simp add: Abs_fresh_iff)
0f0335d91456 solved subgoals for depth and subst function
Christian Urban <urbanc@in.tum.de>
parents: 2653
diff changeset
  2046
apply(simp add: Abs_fresh_iff)
0f0335d91456 solved subgoals for depth and subst function
Christian Urban <urbanc@in.tum.de>
parents: 2653
diff changeset
  2047
apply(simp add: eqvt_at_def)
0f0335d91456 solved subgoals for depth and subst function
Christian Urban <urbanc@in.tum.de>
parents: 2653
diff changeset
  2048
apply(drule test)
0f0335d91456 solved subgoals for depth and subst function
Christian Urban <urbanc@in.tum.de>
parents: 2653
diff changeset
  2049
apply(simp)
0f0335d91456 solved subgoals for depth and subst function
Christian Urban <urbanc@in.tum.de>
parents: 2653
diff changeset
  2050
apply(subst (2) swap_fresh_fresh)
0f0335d91456 solved subgoals for depth and subst function
Christian Urban <urbanc@in.tum.de>
parents: 2653
diff changeset
  2051
apply(simp)
0f0335d91456 solved subgoals for depth and subst function
Christian Urban <urbanc@in.tum.de>
parents: 2653
diff changeset
  2052
apply(simp)
0f0335d91456 solved subgoals for depth and subst function
Christian Urban <urbanc@in.tum.de>
parents: 2653
diff changeset
  2053
apply(subst (2) swap_fresh_fresh)
0f0335d91456 solved subgoals for depth and subst function
Christian Urban <urbanc@in.tum.de>
parents: 2653
diff changeset
  2054
apply(simp)
0f0335d91456 solved subgoals for depth and subst function
Christian Urban <urbanc@in.tum.de>
parents: 2653
diff changeset
  2055
apply(simp)
0f0335d91456 solved subgoals for depth and subst function
Christian Urban <urbanc@in.tum.de>
parents: 2653
diff changeset
  2056
apply(subst (3) swap_fresh_fresh)
0f0335d91456 solved subgoals for depth and subst function
Christian Urban <urbanc@in.tum.de>
parents: 2653
diff changeset
  2057
apply(simp)
0f0335d91456 solved subgoals for depth and subst function
Christian Urban <urbanc@in.tum.de>
parents: 2653
diff changeset
  2058
apply(simp)
0f0335d91456 solved subgoals for depth and subst function
Christian Urban <urbanc@in.tum.de>
parents: 2653
diff changeset
  2059
apply(subst (3) swap_fresh_fresh)
0f0335d91456 solved subgoals for depth and subst function
Christian Urban <urbanc@in.tum.de>
parents: 2653
diff changeset
  2060
apply(simp)
0f0335d91456 solved subgoals for depth and subst function
Christian Urban <urbanc@in.tum.de>
parents: 2653
diff changeset
  2061
apply(simp)
0f0335d91456 solved subgoals for depth and subst function
Christian Urban <urbanc@in.tum.de>
parents: 2653
diff changeset
  2062
apply(simp)
0f0335d91456 solved subgoals for depth and subst function
Christian Urban <urbanc@in.tum.de>
parents: 2653
diff changeset
  2063
apply(simp add: finite_supp)
0f0335d91456 solved subgoals for depth and subst function
Christian Urban <urbanc@in.tum.de>
parents: 2653
diff changeset
  2064
done
2649
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  2065
2653
d0f774d06e6e added eqvt_at premises in function definition - however not proved at the moment
Christian Urban <urbanc@in.tum.de>
parents: 2652
diff changeset
  2066
(* this should not work *)
d0f774d06e6e added eqvt_at premises in function definition - however not proved at the moment
Christian Urban <urbanc@in.tum.de>
parents: 2652
diff changeset
  2067
nominal_primrec 
d0f774d06e6e added eqvt_at premises in function definition - however not proved at the moment
Christian Urban <urbanc@in.tum.de>
parents: 2652
diff changeset
  2068
  bnds :: "lam \<Rightarrow> name set"
d0f774d06e6e added eqvt_at premises in function definition - however not proved at the moment
Christian Urban <urbanc@in.tum.de>
parents: 2652
diff changeset
  2069
where
d0f774d06e6e added eqvt_at premises in function definition - however not proved at the moment
Christian Urban <urbanc@in.tum.de>
parents: 2652
diff changeset
  2070
  "bnds (Var x) = {}"
d0f774d06e6e added eqvt_at premises in function definition - however not proved at the moment
Christian Urban <urbanc@in.tum.de>
parents: 2652
diff changeset
  2071
| "bnds (App t1 t2) = (bnds t1) \<union> (bnds t2)"
d0f774d06e6e added eqvt_at premises in function definition - however not proved at the moment
Christian Urban <urbanc@in.tum.de>
parents: 2652
diff changeset
  2072
| "bnds (Lam x t) = (bnds t) \<union> {x}"
d0f774d06e6e added eqvt_at premises in function definition - however not proved at the moment
Christian Urban <urbanc@in.tum.de>
parents: 2652
diff changeset
  2073
apply(rule_tac y="x" in lam.exhaust)
d0f774d06e6e added eqvt_at premises in function definition - however not proved at the moment
Christian Urban <urbanc@in.tum.de>
parents: 2652
diff changeset
  2074
apply(simp_all)[3]
d0f774d06e6e added eqvt_at premises in function definition - however not proved at the moment
Christian Urban <urbanc@in.tum.de>
parents: 2652
diff changeset
  2075
apply(simp_all only: lam.distinct)
d0f774d06e6e added eqvt_at premises in function definition - however not proved at the moment
Christian Urban <urbanc@in.tum.de>
parents: 2652
diff changeset
  2076
apply(simp add: lam.eq_iff)
d0f774d06e6e added eqvt_at premises in function definition - however not proved at the moment
Christian Urban <urbanc@in.tum.de>
parents: 2652
diff changeset
  2077
apply(simp add: lam.eq_iff)
d0f774d06e6e added eqvt_at premises in function definition - however not proved at the moment
Christian Urban <urbanc@in.tum.de>
parents: 2652
diff changeset
  2078
apply(simp add: lam.eq_iff)
d0f774d06e6e added eqvt_at premises in function definition - however not proved at the moment
Christian Urban <urbanc@in.tum.de>
parents: 2652
diff changeset
  2079
sorry
2649
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  2080
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  2081
end
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  2082
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  2083
a8ebcb368a15 a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  2084