LamEx.thy
author Cezary Kaliszyk <kaliszyk@in.tum.de>
Fri, 30 Oct 2009 18:31:06 +0100
changeset 251 c770f36f9459
parent 249 7dec34d12328
child 252 e30997c88050
permissions -rw-r--r--
Regularization
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
201
1ac36993cc71 added an example about lambda-terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     1
theory LamEx
1ac36993cc71 added an example about lambda-terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     2
imports Nominal QuotMain
1ac36993cc71 added an example about lambda-terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     3
begin
1ac36993cc71 added an example about lambda-terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     4
1ac36993cc71 added an example about lambda-terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     5
atom_decl name
1ac36993cc71 added an example about lambda-terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     6
243
22715cab3995 added fv-function
Christian Urban <urbanc@in.tum.de>
parents: 240
diff changeset
     7
thm abs_fresh(1)
22715cab3995 added fv-function
Christian Urban <urbanc@in.tum.de>
parents: 240
diff changeset
     8
201
1ac36993cc71 added an example about lambda-terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     9
nominal_datatype rlam =
1ac36993cc71 added an example about lambda-terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    10
  rVar "name"
1ac36993cc71 added an example about lambda-terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    11
| rApp "rlam" "rlam"
1ac36993cc71 added an example about lambda-terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    12
| rLam "name" "rlam"
1ac36993cc71 added an example about lambda-terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    13
1ac36993cc71 added an example about lambda-terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    14
243
22715cab3995 added fv-function
Christian Urban <urbanc@in.tum.de>
parents: 240
diff changeset
    15
function
22715cab3995 added fv-function
Christian Urban <urbanc@in.tum.de>
parents: 240
diff changeset
    16
  rfv :: "rlam \<Rightarrow> name set"
22715cab3995 added fv-function
Christian Urban <urbanc@in.tum.de>
parents: 240
diff changeset
    17
where
22715cab3995 added fv-function
Christian Urban <urbanc@in.tum.de>
parents: 240
diff changeset
    18
  rfv_var: "rfv (rVar a) = {a}"
22715cab3995 added fv-function
Christian Urban <urbanc@in.tum.de>
parents: 240
diff changeset
    19
| rfv_app: "rfv (rApp t1 t2) = (rfv t1) \<union> (rfv t2)"
22715cab3995 added fv-function
Christian Urban <urbanc@in.tum.de>
parents: 240
diff changeset
    20
| rfv_lam: "rfv (rLam a t) = (rfv t) - {a}"
22715cab3995 added fv-function
Christian Urban <urbanc@in.tum.de>
parents: 240
diff changeset
    21
sorry
22715cab3995 added fv-function
Christian Urban <urbanc@in.tum.de>
parents: 240
diff changeset
    22
247
e83a6e452843 Lemmas about fv.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 246
diff changeset
    23
termination rfv sorry
243
22715cab3995 added fv-function
Christian Urban <urbanc@in.tum.de>
parents: 240
diff changeset
    24
246
6da7d538e5e0 changed the order of rfv and reformulated a3 with rfv
Christian Urban <urbanc@in.tum.de>
parents: 245
diff changeset
    25
inductive 
6da7d538e5e0 changed the order of rfv and reformulated a3 with rfv
Christian Urban <urbanc@in.tum.de>
parents: 245
diff changeset
    26
  alpha :: "rlam \<Rightarrow> rlam \<Rightarrow> bool" ("_ \<approx> _" [100, 100] 100)
6da7d538e5e0 changed the order of rfv and reformulated a3 with rfv
Christian Urban <urbanc@in.tum.de>
parents: 245
diff changeset
    27
where
6da7d538e5e0 changed the order of rfv and reformulated a3 with rfv
Christian Urban <urbanc@in.tum.de>
parents: 245
diff changeset
    28
  a1: "a = b \<Longrightarrow> (rVar a) \<approx> (rVar b)"
6da7d538e5e0 changed the order of rfv and reformulated a3 with rfv
Christian Urban <urbanc@in.tum.de>
parents: 245
diff changeset
    29
| a2: "\<lbrakk>t1 \<approx> t2; s1 \<approx> s2\<rbrakk> \<Longrightarrow> rApp t1 s1 \<approx> rApp t2 s2"
6da7d538e5e0 changed the order of rfv and reformulated a3 with rfv
Christian Urban <urbanc@in.tum.de>
parents: 245
diff changeset
    30
| a3: "\<lbrakk>t \<approx> ([(a,b)]\<bullet>s); a \<notin> rfv (rLam b t)\<rbrakk> \<Longrightarrow> rLam a t \<approx> rLam b s"
6da7d538e5e0 changed the order of rfv and reformulated a3 with rfv
Christian Urban <urbanc@in.tum.de>
parents: 245
diff changeset
    31
201
1ac36993cc71 added an example about lambda-terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    32
quotient lam = rlam / alpha
1ac36993cc71 added an example about lambda-terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    33
sorry
1ac36993cc71 added an example about lambda-terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    34
203
7384115df9fd added equiv-thm to the quot_info
Christian Urban <urbanc@in.tum.de>
parents: 201
diff changeset
    35
print_quotients
7384115df9fd added equiv-thm to the quot_info
Christian Urban <urbanc@in.tum.de>
parents: 201
diff changeset
    36
229
13f985a93dbc fixed the definition of alpha; this *breaks* some of the experiments
Christian Urban <urbanc@in.tum.de>
parents: 225
diff changeset
    37
quotient_def (for lam)
13f985a93dbc fixed the definition of alpha; this *breaks* some of the experiments
Christian Urban <urbanc@in.tum.de>
parents: 225
diff changeset
    38
  Var :: "name \<Rightarrow> lam"
13f985a93dbc fixed the definition of alpha; this *breaks* some of the experiments
Christian Urban <urbanc@in.tum.de>
parents: 225
diff changeset
    39
where
13f985a93dbc fixed the definition of alpha; this *breaks* some of the experiments
Christian Urban <urbanc@in.tum.de>
parents: 225
diff changeset
    40
  "Var \<equiv> rVar"
13f985a93dbc fixed the definition of alpha; this *breaks* some of the experiments
Christian Urban <urbanc@in.tum.de>
parents: 225
diff changeset
    41
13f985a93dbc fixed the definition of alpha; this *breaks* some of the experiments
Christian Urban <urbanc@in.tum.de>
parents: 225
diff changeset
    42
quotient_def (for lam)
13f985a93dbc fixed the definition of alpha; this *breaks* some of the experiments
Christian Urban <urbanc@in.tum.de>
parents: 225
diff changeset
    43
  App :: "lam \<Rightarrow> lam \<Rightarrow> lam"
13f985a93dbc fixed the definition of alpha; this *breaks* some of the experiments
Christian Urban <urbanc@in.tum.de>
parents: 225
diff changeset
    44
where
13f985a93dbc fixed the definition of alpha; this *breaks* some of the experiments
Christian Urban <urbanc@in.tum.de>
parents: 225
diff changeset
    45
  "App \<equiv> rApp"
13f985a93dbc fixed the definition of alpha; this *breaks* some of the experiments
Christian Urban <urbanc@in.tum.de>
parents: 225
diff changeset
    46
13f985a93dbc fixed the definition of alpha; this *breaks* some of the experiments
Christian Urban <urbanc@in.tum.de>
parents: 225
diff changeset
    47
quotient_def (for lam)
13f985a93dbc fixed the definition of alpha; this *breaks* some of the experiments
Christian Urban <urbanc@in.tum.de>
parents: 225
diff changeset
    48
  Lam :: "name \<Rightarrow> lam \<Rightarrow> lam"
13f985a93dbc fixed the definition of alpha; this *breaks* some of the experiments
Christian Urban <urbanc@in.tum.de>
parents: 225
diff changeset
    49
where
13f985a93dbc fixed the definition of alpha; this *breaks* some of the experiments
Christian Urban <urbanc@in.tum.de>
parents: 225
diff changeset
    50
  "Lam \<equiv> rLam"
201
1ac36993cc71 added an example about lambda-terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    51
218
df05cd030d2f added infrastructure for defining lifted constants
Christian Urban <urbanc@in.tum.de>
parents: 203
diff changeset
    52
thm Var_def
df05cd030d2f added infrastructure for defining lifted constants
Christian Urban <urbanc@in.tum.de>
parents: 203
diff changeset
    53
thm App_def
df05cd030d2f added infrastructure for defining lifted constants
Christian Urban <urbanc@in.tum.de>
parents: 203
diff changeset
    54
thm Lam_def
df05cd030d2f added infrastructure for defining lifted constants
Christian Urban <urbanc@in.tum.de>
parents: 203
diff changeset
    55
243
22715cab3995 added fv-function
Christian Urban <urbanc@in.tum.de>
parents: 240
diff changeset
    56
quotient_def (for lam)
22715cab3995 added fv-function
Christian Urban <urbanc@in.tum.de>
parents: 240
diff changeset
    57
  fv :: "lam \<Rightarrow> name set"
22715cab3995 added fv-function
Christian Urban <urbanc@in.tum.de>
parents: 240
diff changeset
    58
where
22715cab3995 added fv-function
Christian Urban <urbanc@in.tum.de>
parents: 240
diff changeset
    59
  "fv \<equiv> rfv"
22715cab3995 added fv-function
Christian Urban <urbanc@in.tum.de>
parents: 240
diff changeset
    60
22715cab3995 added fv-function
Christian Urban <urbanc@in.tum.de>
parents: 240
diff changeset
    61
thm fv_def
22715cab3995 added fv-function
Christian Urban <urbanc@in.tum.de>
parents: 240
diff changeset
    62
229
13f985a93dbc fixed the definition of alpha; this *breaks* some of the experiments
Christian Urban <urbanc@in.tum.de>
parents: 225
diff changeset
    63
(* definition of overloaded permutation function *)
13f985a93dbc fixed the definition of alpha; this *breaks* some of the experiments
Christian Urban <urbanc@in.tum.de>
parents: 225
diff changeset
    64
(* for the lifted type lam                       *)
13f985a93dbc fixed the definition of alpha; this *breaks* some of the experiments
Christian Urban <urbanc@in.tum.de>
parents: 225
diff changeset
    65
overloading
13f985a93dbc fixed the definition of alpha; this *breaks* some of the experiments
Christian Urban <urbanc@in.tum.de>
parents: 225
diff changeset
    66
  perm_lam    \<equiv> "perm :: 'x prm \<Rightarrow> lam \<Rightarrow> lam"   (unchecked)
13f985a93dbc fixed the definition of alpha; this *breaks* some of the experiments
Christian Urban <urbanc@in.tum.de>
parents: 225
diff changeset
    67
begin
13f985a93dbc fixed the definition of alpha; this *breaks* some of the experiments
Christian Urban <urbanc@in.tum.de>
parents: 225
diff changeset
    68
13f985a93dbc fixed the definition of alpha; this *breaks* some of the experiments
Christian Urban <urbanc@in.tum.de>
parents: 225
diff changeset
    69
quotient_def (for lam)
13f985a93dbc fixed the definition of alpha; this *breaks* some of the experiments
Christian Urban <urbanc@in.tum.de>
parents: 225
diff changeset
    70
  perm_lam :: "'x prm \<Rightarrow> lam \<Rightarrow> lam"
13f985a93dbc fixed the definition of alpha; this *breaks* some of the experiments
Christian Urban <urbanc@in.tum.de>
parents: 225
diff changeset
    71
where
13f985a93dbc fixed the definition of alpha; this *breaks* some of the experiments
Christian Urban <urbanc@in.tum.de>
parents: 225
diff changeset
    72
  "perm_lam \<equiv> (perm::'x prm \<Rightarrow> rlam \<Rightarrow> rlam)"
13f985a93dbc fixed the definition of alpha; this *breaks* some of the experiments
Christian Urban <urbanc@in.tum.de>
parents: 225
diff changeset
    73
13f985a93dbc fixed the definition of alpha; this *breaks* some of the experiments
Christian Urban <urbanc@in.tum.de>
parents: 225
diff changeset
    74
end
13f985a93dbc fixed the definition of alpha; this *breaks* some of the experiments
Christian Urban <urbanc@in.tum.de>
parents: 225
diff changeset
    75
238
e9cc3a3aa5d1 Tried manually lifting real_alpha
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 237
diff changeset
    76
(*quotient_def (for lam)
e9cc3a3aa5d1 Tried manually lifting real_alpha
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 237
diff changeset
    77
  abs_fun_lam :: "'x prm \<Rightarrow> lam \<Rightarrow> lam"
e9cc3a3aa5d1 Tried manually lifting real_alpha
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 237
diff changeset
    78
where
e9cc3a3aa5d1 Tried manually lifting real_alpha
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 237
diff changeset
    79
  "perm_lam \<equiv> (perm::'x prm \<Rightarrow> rlam \<Rightarrow> rlam)"*)
e9cc3a3aa5d1 Tried manually lifting real_alpha
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 237
diff changeset
    80
e9cc3a3aa5d1 Tried manually lifting real_alpha
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 237
diff changeset
    81
229
13f985a93dbc fixed the definition of alpha; this *breaks* some of the experiments
Christian Urban <urbanc@in.tum.de>
parents: 225
diff changeset
    82
thm perm_lam_def
13f985a93dbc fixed the definition of alpha; this *breaks* some of the experiments
Christian Urban <urbanc@in.tum.de>
parents: 225
diff changeset
    83
13f985a93dbc fixed the definition of alpha; this *breaks* some of the experiments
Christian Urban <urbanc@in.tum.de>
parents: 225
diff changeset
    84
(* lemmas that need to lift *)
234
811f17de4031 Lifting of the 3 lemmas in LamEx
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 233
diff changeset
    85
lemma pi_var_com:
229
13f985a93dbc fixed the definition of alpha; this *breaks* some of the experiments
Christian Urban <urbanc@in.tum.de>
parents: 225
diff changeset
    86
  fixes pi::"'x prm"
234
811f17de4031 Lifting of the 3 lemmas in LamEx
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 233
diff changeset
    87
  shows "(pi\<bullet>rVar a) \<approx> rVar (pi\<bullet>a)"
229
13f985a93dbc fixed the definition of alpha; this *breaks* some of the experiments
Christian Urban <urbanc@in.tum.de>
parents: 225
diff changeset
    88
  sorry
13f985a93dbc fixed the definition of alpha; this *breaks* some of the experiments
Christian Urban <urbanc@in.tum.de>
parents: 225
diff changeset
    89
234
811f17de4031 Lifting of the 3 lemmas in LamEx
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 233
diff changeset
    90
lemma pi_app_com:
229
13f985a93dbc fixed the definition of alpha; this *breaks* some of the experiments
Christian Urban <urbanc@in.tum.de>
parents: 225
diff changeset
    91
  fixes pi::"'x prm"
234
811f17de4031 Lifting of the 3 lemmas in LamEx
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 233
diff changeset
    92
  shows "(pi\<bullet>rApp t1 t2) \<approx> rApp (pi\<bullet>t1) (pi\<bullet>t2)"
229
13f985a93dbc fixed the definition of alpha; this *breaks* some of the experiments
Christian Urban <urbanc@in.tum.de>
parents: 225
diff changeset
    93
  sorry
13f985a93dbc fixed the definition of alpha; this *breaks* some of the experiments
Christian Urban <urbanc@in.tum.de>
parents: 225
diff changeset
    94
234
811f17de4031 Lifting of the 3 lemmas in LamEx
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 233
diff changeset
    95
lemma pi_lam_com:
229
13f985a93dbc fixed the definition of alpha; this *breaks* some of the experiments
Christian Urban <urbanc@in.tum.de>
parents: 225
diff changeset
    96
  fixes pi::"'x prm"
234
811f17de4031 Lifting of the 3 lemmas in LamEx
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 233
diff changeset
    97
  shows "(pi\<bullet>rLam a t) \<approx> rLam (pi\<bullet>a) (pi\<bullet>t)"
229
13f985a93dbc fixed the definition of alpha; this *breaks* some of the experiments
Christian Urban <urbanc@in.tum.de>
parents: 225
diff changeset
    98
  sorry
13f985a93dbc fixed the definition of alpha; this *breaks* some of the experiments
Christian Urban <urbanc@in.tum.de>
parents: 225
diff changeset
    99
243
22715cab3995 added fv-function
Christian Urban <urbanc@in.tum.de>
parents: 240
diff changeset
   100
lemma fv_var:
22715cab3995 added fv-function
Christian Urban <urbanc@in.tum.de>
parents: 240
diff changeset
   101
  shows "fv (Var a) = {a}"
22715cab3995 added fv-function
Christian Urban <urbanc@in.tum.de>
parents: 240
diff changeset
   102
sorry
22715cab3995 added fv-function
Christian Urban <urbanc@in.tum.de>
parents: 240
diff changeset
   103
22715cab3995 added fv-function
Christian Urban <urbanc@in.tum.de>
parents: 240
diff changeset
   104
lemma fv_app:
22715cab3995 added fv-function
Christian Urban <urbanc@in.tum.de>
parents: 240
diff changeset
   105
  shows "fv (App t1 t2) = (fv t1) \<union> (fv t2)"
22715cab3995 added fv-function
Christian Urban <urbanc@in.tum.de>
parents: 240
diff changeset
   106
sorry
22715cab3995 added fv-function
Christian Urban <urbanc@in.tum.de>
parents: 240
diff changeset
   107
22715cab3995 added fv-function
Christian Urban <urbanc@in.tum.de>
parents: 240
diff changeset
   108
lemma fv_lam:
22715cab3995 added fv-function
Christian Urban <urbanc@in.tum.de>
parents: 240
diff changeset
   109
  shows "fv (Lam a t) = (fv t) - {a}"
22715cab3995 added fv-function
Christian Urban <urbanc@in.tum.de>
parents: 240
diff changeset
   110
sorry 
22715cab3995 added fv-function
Christian Urban <urbanc@in.tum.de>
parents: 240
diff changeset
   111
201
1ac36993cc71 added an example about lambda-terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   112
lemma real_alpha:
242
47de63a883c2 The proper real_alpha
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 240
diff changeset
   113
  assumes "t = [(a,b)]\<bullet>s" "a\<sharp>[b].s"
201
1ac36993cc71 added an example about lambda-terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   114
  shows "Lam a t = Lam b s"
217
9cc87d02190a First experiments with Lambda
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 203
diff changeset
   115
sorry
9cc87d02190a First experiments with Lambda
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 203
diff changeset
   116
9cc87d02190a First experiments with Lambda
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 203
diff changeset
   117
9cc87d02190a First experiments with Lambda
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 203
diff changeset
   118
9cc87d02190a First experiments with Lambda
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 203
diff changeset
   119
9cc87d02190a First experiments with Lambda
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 203
diff changeset
   120
9cc87d02190a First experiments with Lambda
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 203
diff changeset
   121
(* Construction Site code *)
9cc87d02190a First experiments with Lambda
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 203
diff changeset
   122
234
811f17de4031 Lifting of the 3 lemmas in LamEx
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 233
diff changeset
   123
lemma perm_rsp: "(op = ===> alpha ===> alpha) op \<bullet> op \<bullet>"
229
13f985a93dbc fixed the definition of alpha; this *breaks* some of the experiments
Christian Urban <urbanc@in.tum.de>
parents: 225
diff changeset
   124
  apply(auto)
13f985a93dbc fixed the definition of alpha; this *breaks* some of the experiments
Christian Urban <urbanc@in.tum.de>
parents: 225
diff changeset
   125
  (* this is propably true if some type conditions are imposed ;o) *)
13f985a93dbc fixed the definition of alpha; this *breaks* some of the experiments
Christian Urban <urbanc@in.tum.de>
parents: 225
diff changeset
   126
  sorry
13f985a93dbc fixed the definition of alpha; this *breaks* some of the experiments
Christian Urban <urbanc@in.tum.de>
parents: 225
diff changeset
   127
234
811f17de4031 Lifting of the 3 lemmas in LamEx
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 233
diff changeset
   128
lemma fresh_rsp: "(op = ===> alpha ===> op =) fresh fresh" 
229
13f985a93dbc fixed the definition of alpha; this *breaks* some of the experiments
Christian Urban <urbanc@in.tum.de>
parents: 225
diff changeset
   129
  apply(auto)
13f985a93dbc fixed the definition of alpha; this *breaks* some of the experiments
Christian Urban <urbanc@in.tum.de>
parents: 225
diff changeset
   130
  (* this is probably only true if some type conditions are imposed *)
13f985a93dbc fixed the definition of alpha; this *breaks* some of the experiments
Christian Urban <urbanc@in.tum.de>
parents: 225
diff changeset
   131
  sorry
13f985a93dbc fixed the definition of alpha; this *breaks* some of the experiments
Christian Urban <urbanc@in.tum.de>
parents: 225
diff changeset
   132
234
811f17de4031 Lifting of the 3 lemmas in LamEx
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 233
diff changeset
   133
lemma rVar_rsp: "(op = ===> alpha) rVar rVar"
811f17de4031 Lifting of the 3 lemmas in LamEx
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 233
diff changeset
   134
  apply(auto)
811f17de4031 Lifting of the 3 lemmas in LamEx
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 233
diff changeset
   135
  apply(rule a1)
811f17de4031 Lifting of the 3 lemmas in LamEx
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 233
diff changeset
   136
  apply(simp)
811f17de4031 Lifting of the 3 lemmas in LamEx
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 233
diff changeset
   137
  done
811f17de4031 Lifting of the 3 lemmas in LamEx
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 233
diff changeset
   138
811f17de4031 Lifting of the 3 lemmas in LamEx
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 233
diff changeset
   139
lemma rApp_rsp: "(alpha ===> alpha ===> alpha) rApp rApp"
811f17de4031 Lifting of the 3 lemmas in LamEx
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 233
diff changeset
   140
  apply(auto)
811f17de4031 Lifting of the 3 lemmas in LamEx
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 233
diff changeset
   141
  apply(rule a2)
811f17de4031 Lifting of the 3 lemmas in LamEx
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 233
diff changeset
   142
  apply (assumption)
811f17de4031 Lifting of the 3 lemmas in LamEx
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 233
diff changeset
   143
  apply (assumption)
811f17de4031 Lifting of the 3 lemmas in LamEx
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 233
diff changeset
   144
  done
811f17de4031 Lifting of the 3 lemmas in LamEx
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 233
diff changeset
   145
811f17de4031 Lifting of the 3 lemmas in LamEx
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 233
diff changeset
   146
lemma rLam_rsp: "(op = ===> alpha ===> alpha) rLam rLam"
229
13f985a93dbc fixed the definition of alpha; this *breaks* some of the experiments
Christian Urban <urbanc@in.tum.de>
parents: 225
diff changeset
   147
  apply(auto)
13f985a93dbc fixed the definition of alpha; this *breaks* some of the experiments
Christian Urban <urbanc@in.tum.de>
parents: 225
diff changeset
   148
  apply(rule a3)
13f985a93dbc fixed the definition of alpha; this *breaks* some of the experiments
Christian Urban <urbanc@in.tum.de>
parents: 225
diff changeset
   149
  apply(rule_tac t="[(x,x)]\<bullet>y" and s="y" in subst)
13f985a93dbc fixed the definition of alpha; this *breaks* some of the experiments
Christian Urban <urbanc@in.tum.de>
parents: 225
diff changeset
   150
  apply(rule sym)
13f985a93dbc fixed the definition of alpha; this *breaks* some of the experiments
Christian Urban <urbanc@in.tum.de>
parents: 225
diff changeset
   151
  apply(rule trans)
13f985a93dbc fixed the definition of alpha; this *breaks* some of the experiments
Christian Urban <urbanc@in.tum.de>
parents: 225
diff changeset
   152
  apply(rule pt_name3)
13f985a93dbc fixed the definition of alpha; this *breaks* some of the experiments
Christian Urban <urbanc@in.tum.de>
parents: 225
diff changeset
   153
  apply(rule at_ds1[OF at_name_inst])
13f985a93dbc fixed the definition of alpha; this *breaks* some of the experiments
Christian Urban <urbanc@in.tum.de>
parents: 225
diff changeset
   154
  apply(simp add: pt_name1)
13f985a93dbc fixed the definition of alpha; this *breaks* some of the experiments
Christian Urban <urbanc@in.tum.de>
parents: 225
diff changeset
   155
  apply(assumption)
13f985a93dbc fixed the definition of alpha; this *breaks* some of the experiments
Christian Urban <urbanc@in.tum.de>
parents: 225
diff changeset
   156
  apply(simp add: abs_fresh)
13f985a93dbc fixed the definition of alpha; this *breaks* some of the experiments
Christian Urban <urbanc@in.tum.de>
parents: 225
diff changeset
   157
  done
217
9cc87d02190a First experiments with Lambda
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 203
diff changeset
   158
247
e83a6e452843 Lemmas about fv.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 246
diff changeset
   159
lemma rfv_rsp: "(alpha ===> op =) rfv rfv"
e83a6e452843 Lemmas about fv.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 246
diff changeset
   160
  sorry
217
9cc87d02190a First experiments with Lambda
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 203
diff changeset
   161
9cc87d02190a First experiments with Lambda
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 203
diff changeset
   162
ML {* val qty = @{typ "lam"} *}
240
6cff34032a00 Cleaning also in Lam
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 238
diff changeset
   163
ML {* val (rty, rel, rel_refl, rel_eqv) = lookup_quot_data @{context} qty *}
247
e83a6e452843 Lemmas about fv.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 246
diff changeset
   164
ML {* val defs = @{thms Var_def App_def Lam_def perm_lam_def fv_def} *}
240
6cff34032a00 Cleaning also in Lam
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 238
diff changeset
   165
ML {* val consts = lookup_quot_consts defs *}
6cff34032a00 Cleaning also in Lam
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 238
diff changeset
   166
ML {* val (trans2, reps_same, quot) = lookup_quot_thms @{context} "lam" *}
247
e83a6e452843 Lemmas about fv.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 246
diff changeset
   167
ML {* val rsp_thms = @{thms perm_rsp fresh_rsp rVar_rsp rApp_rsp rLam_rsp rfv_rsp} @ @{thms ho_all_prs ho_ex_prs} *}
240
6cff34032a00 Cleaning also in Lam
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 238
diff changeset
   168
ML {* fun lift_thm_lam lthy t = lift_thm lthy qty "lam" rsp_thms defs t *}
237
80f1df49b940 More tests in Lam
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 234
diff changeset
   169
249
7dec34d12328 added some facts about fresh and support of lam
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
   170
ML {* val pi_var = lift_thm_lam @{context} @{thm pi_var_com} *}
7dec34d12328 added some facts about fresh and support of lam
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
   171
ML {* val pi_app = lift_thm_lam @{context} @{thm pi_app_com} *}
7dec34d12328 added some facts about fresh and support of lam
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
   172
ML {* val pi_lam = lift_thm_lam @{context} @{thm pi_lam_com} *}
7dec34d12328 added some facts about fresh and support of lam
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
   173
7dec34d12328 added some facts about fresh and support of lam
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
   174
ML {* val fv_var = lift_thm_lam @{context} @{thm rfv_var} *}
7dec34d12328 added some facts about fresh and support of lam
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
   175
ML {* val fv_app = lift_thm_lam @{context} @{thm rfv_app} *}
7dec34d12328 added some facts about fresh and support of lam
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
   176
ML {* val fv_lam = lift_thm_lam @{context} @{thm rfv_lam} *}
7dec34d12328 added some facts about fresh and support of lam
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
   177
7dec34d12328 added some facts about fresh and support of lam
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
   178
ML {* val a1 = lift_thm_lam @{context} @{thm a1} *}
7dec34d12328 added some facts about fresh and support of lam
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
   179
ML {* val a2 = lift_thm_lam @{context} @{thm a1} *}
7dec34d12328 added some facts about fresh and support of lam
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
   180
ML {* val a3 = lift_thm_lam @{context} @{thm a3} *}
7dec34d12328 added some facts about fresh and support of lam
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
   181
7dec34d12328 added some facts about fresh and support of lam
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
   182
local_setup {*
7dec34d12328 added some facts about fresh and support of lam
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
   183
  Quotient.note (@{binding "pi_var"}, pi_var) #> snd #>
7dec34d12328 added some facts about fresh and support of lam
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
   184
  Quotient.note (@{binding "pi_app"}, pi_app) #> snd #>
7dec34d12328 added some facts about fresh and support of lam
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
   185
  Quotient.note (@{binding "pi_lam"}, pi_lam) #> snd #>
7dec34d12328 added some facts about fresh and support of lam
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
   186
  Quotient.note (@{binding "a1"}, a1) #> snd #>
7dec34d12328 added some facts about fresh and support of lam
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
   187
  Quotient.note (@{binding "a2"}, a2) #> snd #>
7dec34d12328 added some facts about fresh and support of lam
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
   188
  Quotient.note (@{binding "a3"}, a3) #> snd
7dec34d12328 added some facts about fresh and support of lam
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
   189
*}
237
80f1df49b940 More tests in Lam
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 234
diff changeset
   190
249
7dec34d12328 added some facts about fresh and support of lam
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
   191
thm alpha.cases
7dec34d12328 added some facts about fresh and support of lam
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
   192
7dec34d12328 added some facts about fresh and support of lam
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
   193
thm rlam.inject
7dec34d12328 added some facts about fresh and support of lam
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
   194
thm pi_var
7dec34d12328 added some facts about fresh and support of lam
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
   195
7dec34d12328 added some facts about fresh and support of lam
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
   196
lemma var_inject:
7dec34d12328 added some facts about fresh and support of lam
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
   197
  shows "(Var a = Var b) = (a = b)"
7dec34d12328 added some facts about fresh and support of lam
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
   198
sorry
247
e83a6e452843 Lemmas about fv.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 246
diff changeset
   199
249
7dec34d12328 added some facts about fresh and support of lam
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
   200
lemma var_supp:
7dec34d12328 added some facts about fresh and support of lam
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
   201
  shows "supp (Var a) = ((supp a)::name set)"
7dec34d12328 added some facts about fresh and support of lam
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
   202
  apply(simp add: supp_def)
7dec34d12328 added some facts about fresh and support of lam
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
   203
  apply(simp add: pi_var)
7dec34d12328 added some facts about fresh and support of lam
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
   204
  apply(simp add: var_inject)
7dec34d12328 added some facts about fresh and support of lam
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
   205
  done
7dec34d12328 added some facts about fresh and support of lam
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
   206
7dec34d12328 added some facts about fresh and support of lam
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
   207
lemma var_fresh:
7dec34d12328 added some facts about fresh and support of lam
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
   208
  fixes a::"name"
7dec34d12328 added some facts about fresh and support of lam
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
   209
  shows "(a\<sharp>(Var b)) = (a\<sharp>b)"
7dec34d12328 added some facts about fresh and support of lam
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
   210
  apply(simp add: fresh_def)
7dec34d12328 added some facts about fresh and support of lam
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
   211
  apply(simp add: var_supp)
7dec34d12328 added some facts about fresh and support of lam
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
   212
  done
247
e83a6e452843 Lemmas about fv.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 246
diff changeset
   213
e83a6e452843 Lemmas about fv.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 246
diff changeset
   214
e83a6e452843 Lemmas about fv.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 246
diff changeset
   215
e83a6e452843 Lemmas about fv.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 246
diff changeset
   216
e83a6e452843 Lemmas about fv.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 246
diff changeset
   217
e83a6e452843 Lemmas about fv.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 246
diff changeset
   218
e83a6e452843 Lemmas about fv.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 246
diff changeset
   219
e83a6e452843 Lemmas about fv.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 246
diff changeset
   220
251
c770f36f9459 Regularization
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 249
diff changeset
   221
lemma get_rid: "(\<And>x. (P x \<longrightarrow> Q x)) \<Longrightarrow> ((\<forall>x. P x) \<longrightarrow> (\<forall>x. Q x))"
c770f36f9459 Regularization
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 249
diff changeset
   222
apply (auto)
c770f36f9459 Regularization
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 249
diff changeset
   223
done
c770f36f9459 Regularization
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 249
diff changeset
   224
c770f36f9459 Regularization
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 249
diff changeset
   225
lemma get_rid2: "(c \<longrightarrow> a) \<Longrightarrow> (a \<Longrightarrow> b \<longrightarrow> d) \<Longrightarrow> (a \<longrightarrow> b) \<longrightarrow> (c \<longrightarrow> d)"
c770f36f9459 Regularization
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 249
diff changeset
   226
apply (auto)
c770f36f9459 Regularization
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 249
diff changeset
   227
done
c770f36f9459 Regularization
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 249
diff changeset
   228
c770f36f9459 Regularization
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 249
diff changeset
   229
ML {* val t_a = atomize_thm @{thm alpha.cases} *}
c770f36f9459 Regularization
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 249
diff changeset
   230
prove {* build_regularize_goal t_a rty rel @{context}  *}
c770f36f9459 Regularization
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 249
diff changeset
   231
  ML_prf {*  fun tac ctxt =
c770f36f9459 Regularization
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 249
diff changeset
   232
     (FIRST' [
c770f36f9459 Regularization
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 249
diff changeset
   233
      rtac rel_refl,
c770f36f9459 Regularization
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 249
diff changeset
   234
      atac,
c770f36f9459 Regularization
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 249
diff changeset
   235
      rtac @{thm get_rid},
c770f36f9459 Regularization
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 249
diff changeset
   236
      rtac @{thm get_rid2},
c770f36f9459 Regularization
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 249
diff changeset
   237
      (fn i => CHANGED (asm_full_simp_tac ((Simplifier.context ctxt HOL_ss) addsimps
c770f36f9459 Regularization
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 249
diff changeset
   238
        [(@{thm equiv_res_forall} OF [rel_eqv]),
c770f36f9459 Regularization
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 249
diff changeset
   239
         (@{thm equiv_res_exists} OF [rel_eqv])]) i)),
c770f36f9459 Regularization
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 249
diff changeset
   240
      (rtac @{thm impI} THEN' (asm_full_simp_tac (Simplifier.context ctxt HOL_ss)) THEN' rtac rel_refl)
c770f36f9459 Regularization
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 249
diff changeset
   241
    ]);
c770f36f9459 Regularization
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 249
diff changeset
   242
 *}
c770f36f9459 Regularization
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 249
diff changeset
   243
  apply (atomize(full))
c770f36f9459 Regularization
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 249
diff changeset
   244
  apply (tactic {* tac @{context} 1 *})
c770f36f9459 Regularization
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 249
diff changeset
   245
  apply (rule get_rid)
c770f36f9459 Regularization
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 249
diff changeset
   246
  apply (rule get_rid)
c770f36f9459 Regularization
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 249
diff changeset
   247
  apply (rule get_rid2)
c770f36f9459 Regularization
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 249
diff changeset
   248
  apply (simp)
c770f36f9459 Regularization
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 249
diff changeset
   249
  apply (rule get_rid)
c770f36f9459 Regularization
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 249
diff changeset
   250
  apply (rule get_rid2)
c770f36f9459 Regularization
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 249
diff changeset
   251
  apply (rule get_rid)
c770f36f9459 Regularization
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 249
diff changeset
   252
  apply (rule get_rid2)
c770f36f9459 Regularization
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 249
diff changeset
   253
  apply (rule impI)
c770f36f9459 Regularization
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 249
diff changeset
   254
  apply (simp)
c770f36f9459 Regularization
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 249
diff changeset
   255
  apply (tactic {* tac @{context} 1 *})
c770f36f9459 Regularization
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 249
diff changeset
   256
  apply (rule get_rid2)
c770f36f9459 Regularization
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 249
diff changeset
   257
  apply (rule impI)
c770f36f9459 Regularization
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 249
diff changeset
   258
  apply (simp)
c770f36f9459 Regularization
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 249
diff changeset
   259
  apply (tactic {* tac @{context} 1 *})
c770f36f9459 Regularization
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 249
diff changeset
   260
  apply (simp)
c770f36f9459 Regularization
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 249
diff changeset
   261
  apply (rule get_rid2)
c770f36f9459 Regularization
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 249
diff changeset
   262
  apply (rule get_rid)
c770f36f9459 Regularization
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 249
diff changeset
   263
  apply (rule get_rid)
c770f36f9459 Regularization
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 249
diff changeset
   264
  apply (rule get_rid)
c770f36f9459 Regularization
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 249
diff changeset
   265
  apply (rule get_rid2)
c770f36f9459 Regularization
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 249
diff changeset
   266
  apply (rule impI)
c770f36f9459 Regularization
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 249
diff changeset
   267
  apply (simp)
c770f36f9459 Regularization
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 249
diff changeset
   268
  apply (tactic {* tac @{context} 1 *})
c770f36f9459 Regularization
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 249
diff changeset
   269
  apply (rule get_rid)
c770f36f9459 Regularization
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 249
diff changeset
   270
  apply (rule get_rid2)
c770f36f9459 Regularization
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 249
diff changeset
   271
c770f36f9459 Regularization
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 249
diff changeset
   272
c770f36f9459 Regularization
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 249
diff changeset
   273
  apply (tactic {* compose_tac (false, @{thm get_rid}, 1) 1 *})
c770f36f9459 Regularization
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 249
diff changeset
   274
ML_prf {*
c770f36f9459 Regularization
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 249
diff changeset
   275
fun focus_compose t = Subgoal.FOCUS (fn {concl, ...} =>
c770f36f9459 Regularization
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 249
diff changeset
   276
  rtac t 1)
c770f36f9459 Regularization
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 249
diff changeset
   277
*}
c770f36f9459 Regularization
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 249
diff changeset
   278
thm spec[of _ "x"]
c770f36f9459 Regularization
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 249
diff changeset
   279
  apply (rule allI)
c770f36f9459 Regularization
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 249
diff changeset
   280
  apply (drule_tac x="a2" in spec)
c770f36f9459 Regularization
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 249
diff changeset
   281
  apply (rule impI)
c770f36f9459 Regularization
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 249
diff changeset
   282
  apply (erule impE)
c770f36f9459 Regularization
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 249
diff changeset
   283
  apply (assumption)
c770f36f9459 Regularization
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 249
diff changeset
   284
  apply (rule allI)
c770f36f9459 Regularization
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 249
diff changeset
   285
  apply (drule_tac x="P" in spec)
c770f36f9459 Regularization
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 249
diff changeset
   286
  apply (atomize(full))
c770f36f9459 Regularization
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 249
diff changeset
   287
  apply (rule allI)
c770f36f9459 Regularization
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 249
diff changeset
   288
  apply (rule allI)
c770f36f9459 Regularization
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 249
diff changeset
   289
  apply (rule allI)
c770f36f9459 Regularization
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 249
diff changeset
   290
  apply (rule impI)
c770f36f9459 Regularization
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 249
diff changeset
   291
  apply (rule get_rid2)
c770f36f9459 Regularization
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 249
diff changeset
   292
c770f36f9459 Regularization
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 249
diff changeset
   293
  thm get_rid2
c770f36f9459 Regularization
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 249
diff changeset
   294
  apply (tactic {* compose_tac (false, @{thm get_rid2}, 0) 1 *})
c770f36f9459 Regularization
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 249
diff changeset
   295
c770f36f9459 Regularization
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 249
diff changeset
   296
  thm spec
c770f36f9459 Regularization
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 249
diff changeset
   297
c770f36f9459 Regularization
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 249
diff changeset
   298
  ML_prf {* val t = snd (Subgoal.focus @{context} 1 (Isar.goal ())) *}
c770f36f9459 Regularization
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 249
diff changeset
   299
ML_prf {* Seq.pull (compose_tac (false, @{thm get_rid}, 2) 1 t) *}
c770f36f9459 Regularization
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 249
diff changeset
   300
c770f36f9459 Regularization
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 249
diff changeset
   301
c770f36f9459 Regularization
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 249
diff changeset
   302
  thm get_rid
c770f36f9459 Regularization
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 249
diff changeset
   303
  apply (rule allI)
c770f36f9459 Regularization
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 249
diff changeset
   304
  apply (drule spec)
c770f36f9459 Regularization
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 249
diff changeset
   305
c770f36f9459 Regularization
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 249
diff changeset
   306
  thm spec
c770f36f9459 Regularization
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 249
diff changeset
   307
  apply (tactic {* compose_tac (false, @{thm get_rid}, 1) 1 *})
c770f36f9459 Regularization
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 249
diff changeset
   308
c770f36f9459 Regularization
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 249
diff changeset
   309
  apply (tactic {* tac @{context} 1 *})
c770f36f9459 Regularization
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 249
diff changeset
   310
  apply (tactic {* tac @{context} 1 *})
c770f36f9459 Regularization
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 249
diff changeset
   311
  apply (rule impI)
c770f36f9459 Regularization
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 249
diff changeset
   312
  apply (erule impE)
c770f36f9459 Regularization
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 249
diff changeset
   313
  apply (assumption)
c770f36f9459 Regularization
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 249
diff changeset
   314
  apply (tactic {* tac @{context} 1 *})
c770f36f9459 Regularization
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 249
diff changeset
   315
  apply (rule impI)
c770f36f9459 Regularization
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 249
diff changeset
   316
  apply (erule impE)
c770f36f9459 Regularization
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 249
diff changeset
   317
  apply (tactic {* tac @{context} 1 *})
c770f36f9459 Regularization
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 249
diff changeset
   318
  apply (tactic {* tac @{context} 1 *})
c770f36f9459 Regularization
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 249
diff changeset
   319
  apply (tactic {* tac @{context} 1 *})
c770f36f9459 Regularization
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 249
diff changeset
   320
  apply (tactic {* tac @{context} 1 *})
c770f36f9459 Regularization
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 249
diff changeset
   321
c770f36f9459 Regularization
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 249
diff changeset
   322
ML {* val t_r = regularize t_a rty rel rel_eqv rel_refl @{context} *}
c770f36f9459 Regularization
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 249
diff changeset
   323
ML {* val t_t = repabs @{context} t_r consts rty qty quot rel_refl trans2 rsp_thms *}
c770f36f9459 Regularization
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 249
diff changeset
   324
ML {* val t_l = repeat_eqsubst_thm @{context} simp_lam_prs_thms t_t *}
c770f36f9459 Regularization
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 249
diff changeset
   325
ML {* val t_a = simp_allex_prs @{context} quot t_l *}
c770f36f9459 Regularization
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 249
diff changeset
   326
ML {* val t_d = repeat_eqsubst_thm @{context} t_defs_sym
247
e83a6e452843 Lemmas about fv.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 246
diff changeset
   327
e83a6e452843 Lemmas about fv.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 246
diff changeset
   328
e83a6e452843 Lemmas about fv.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 246
diff changeset
   329
e83a6e452843 Lemmas about fv.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 246
diff changeset
   330
237
80f1df49b940 More tests in Lam
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 234
diff changeset
   331
fun
80f1df49b940 More tests in Lam
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 234
diff changeset
   332
  option_map::"('a \<Rightarrow> 'b) \<Rightarrow> ('a noption) \<Rightarrow> ('b noption)"
80f1df49b940 More tests in Lam
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 234
diff changeset
   333
where
80f1df49b940 More tests in Lam
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 234
diff changeset
   334
  "option_map f (nSome x) = nSome (f x)"
80f1df49b940 More tests in Lam
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 234
diff changeset
   335
| "option_map f nNone = nNone"
80f1df49b940 More tests in Lam
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 234
diff changeset
   336
80f1df49b940 More tests in Lam
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 234
diff changeset
   337
fun
80f1df49b940 More tests in Lam
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 234
diff changeset
   338
  option_rel
80f1df49b940 More tests in Lam
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 234
diff changeset
   339
where
80f1df49b940 More tests in Lam
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 234
diff changeset
   340
  "option_rel r (nSome x) (nSome y) = r x y"
80f1df49b940 More tests in Lam
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 234
diff changeset
   341
| "option_rel r _ _ = False"
80f1df49b940 More tests in Lam
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 234
diff changeset
   342
80f1df49b940 More tests in Lam
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 234
diff changeset
   343
declare [[map noption = (option_map, option_rel)]]
80f1df49b940 More tests in Lam
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 234
diff changeset
   344
80f1df49b940 More tests in Lam
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 234
diff changeset
   345
lemma OPT_QUOTIENT:
80f1df49b940 More tests in Lam
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 234
diff changeset
   346
  assumes q: "QUOTIENT R Abs Rep"
80f1df49b940 More tests in Lam
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 234
diff changeset
   347
  shows "QUOTIENT (option_rel R) (option_map Abs) (option_map Rep)"
80f1df49b940 More tests in Lam
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 234
diff changeset
   348
  apply (unfold QUOTIENT_def)
80f1df49b940 More tests in Lam
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 234
diff changeset
   349
  apply (auto)
80f1df49b940 More tests in Lam
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 234
diff changeset
   350
  using q
80f1df49b940 More tests in Lam
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 234
diff changeset
   351
  apply (unfold QUOTIENT_def)
80f1df49b940 More tests in Lam
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 234
diff changeset
   352
  apply (case_tac "a :: 'b noption")
80f1df49b940 More tests in Lam
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 234
diff changeset
   353
  apply (simp)
80f1df49b940 More tests in Lam
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 234
diff changeset
   354
  apply (simp)
80f1df49b940 More tests in Lam
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 234
diff changeset
   355
  apply (case_tac "a :: 'b noption")
80f1df49b940 More tests in Lam
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 234
diff changeset
   356
  apply (simp only: option_map.simps)
80f1df49b940 More tests in Lam
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 234
diff changeset
   357
  apply (subst option_rel.simps)
80f1df49b940 More tests in Lam
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 234
diff changeset
   358
  (* Simp starts hanging so don't know how to continue *)
80f1df49b940 More tests in Lam
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 234
diff changeset
   359
  sorry
80f1df49b940 More tests in Lam
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 234
diff changeset
   360
240
6cff34032a00 Cleaning also in Lam
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 238
diff changeset
   361
(* Not sure if it make sense or if it will be needed *)
237
80f1df49b940 More tests in Lam
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 234
diff changeset
   362
lemma abs_fun_rsp: "(op = ===> alpha ===> op = ===> op =) abs_fun abs_fun"
80f1df49b940 More tests in Lam
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 234
diff changeset
   363
sorry
80f1df49b940 More tests in Lam
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 234
diff changeset
   364
80f1df49b940 More tests in Lam
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 234
diff changeset
   365
(* Should not be needed *)
80f1df49b940 More tests in Lam
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 234
diff changeset
   366
lemma eq_rsp2: "((op = ===> op =) ===> (op = ===> op =) ===> op =) op = op ="
80f1df49b940 More tests in Lam
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 234
diff changeset
   367
apply auto
80f1df49b940 More tests in Lam
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 234
diff changeset
   368
apply (rule ext)
80f1df49b940 More tests in Lam
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 234
diff changeset
   369
apply auto
80f1df49b940 More tests in Lam
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 234
diff changeset
   370
apply (rule ext)
80f1df49b940 More tests in Lam
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 234
diff changeset
   371
apply auto
80f1df49b940 More tests in Lam
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 234
diff changeset
   372
done
217
9cc87d02190a First experiments with Lambda
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 203
diff changeset
   373
237
80f1df49b940 More tests in Lam
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 234
diff changeset
   374
(* Should not be needed *)
80f1df49b940 More tests in Lam
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 234
diff changeset
   375
lemma perm_rsp_eq: "(op = ===> (op = ===> op =) ===> op = ===> op =) op \<bullet> op \<bullet>"
80f1df49b940 More tests in Lam
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 234
diff changeset
   376
apply auto
80f1df49b940 More tests in Lam
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 234
diff changeset
   377
thm arg_cong2
80f1df49b940 More tests in Lam
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 234
diff changeset
   378
apply (rule_tac f="perm x" in arg_cong2)
80f1df49b940 More tests in Lam
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 234
diff changeset
   379
apply (auto)
80f1df49b940 More tests in Lam
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 234
diff changeset
   380
apply (rule ext)
80f1df49b940 More tests in Lam
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 234
diff changeset
   381
apply (auto)
80f1df49b940 More tests in Lam
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 234
diff changeset
   382
done
80f1df49b940 More tests in Lam
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 234
diff changeset
   383
80f1df49b940 More tests in Lam
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 234
diff changeset
   384
(* Should not be needed *)
80f1df49b940 More tests in Lam
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 234
diff changeset
   385
lemma fresh_rsp_eq: "(op = ===> (op = ===> op =) ===> op =) fresh fresh"
80f1df49b940 More tests in Lam
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 234
diff changeset
   386
apply (simp add: FUN_REL.simps)
80f1df49b940 More tests in Lam
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 234
diff changeset
   387
apply (metis ext)
80f1df49b940 More tests in Lam
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 234
diff changeset
   388
done
80f1df49b940 More tests in Lam
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 234
diff changeset
   389
80f1df49b940 More tests in Lam
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 234
diff changeset
   390
(* It is just a test, it doesn't seem true... *)
80f1df49b940 More tests in Lam
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 234
diff changeset
   391
lemma quotient_cheat: "QUOTIENT op = (option_map ABS_lam) (option_map REP_lam)"
80f1df49b940 More tests in Lam
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 234
diff changeset
   392
sorry
80f1df49b940 More tests in Lam
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 234
diff changeset
   393
80f1df49b940 More tests in Lam
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 234
diff changeset
   394
ML {* val rsp_thms = @{thms abs_fun_rsp OPT_QUOTIENT eq_rsp2 quotient_cheat perm_rsp_eq fresh_rsp_eq} @ rsp_thms *}
240
6cff34032a00 Cleaning also in Lam
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 238
diff changeset
   395
ML {* fun lift_thm_lam lthy t = lift_thm lthy qty "lam" rsp_thms defs t *}
237
80f1df49b940 More tests in Lam
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 234
diff changeset
   396
80f1df49b940 More tests in Lam
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 234
diff changeset
   397
thm a3
80f1df49b940 More tests in Lam
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 234
diff changeset
   398
ML {* Toplevel.program (fn () => lift_thm_lam @{context} @{thm a3}) *}
238
e9cc3a3aa5d1 Tried manually lifting real_alpha
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 237
diff changeset
   399
thm a3
e9cc3a3aa5d1 Tried manually lifting real_alpha
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 237
diff changeset
   400
ML {* val t_u1 = eqsubst_thm @{context} @{thms abs_fresh(1)} (atomize_thm @{thm a3}) *}
e9cc3a3aa5d1 Tried manually lifting real_alpha
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 237
diff changeset
   401
ML {* val t_u = MetaSimplifier.rewrite_rule @{thms fresh_def supp_def} t_u1 *}
237
80f1df49b940 More tests in Lam
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 234
diff changeset
   402
247
e83a6e452843 Lemmas about fv.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 246
diff changeset
   403
(* T_U *)
e83a6e452843 Lemmas about fv.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 246
diff changeset
   404
e83a6e452843 Lemmas about fv.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 246
diff changeset
   405
ML {* val t_a = atomize_thm @{thm rfv_var} *}
237
80f1df49b940 More tests in Lam
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 234
diff changeset
   406
ML {* val t_r = regularize t_a rty rel rel_eqv @{context} *}
247
e83a6e452843 Lemmas about fv.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 246
diff changeset
   407
ML {* val t_t = repabs @{context} t_r consts rty qty quot rel_refl trans2 rsp_thms *}
e83a6e452843 Lemmas about fv.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 246
diff changeset
   408
238
e9cc3a3aa5d1 Tried manually lifting real_alpha
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 237
diff changeset
   409
ML {* fun r_mk_comb_tac_lam ctxt =
e9cc3a3aa5d1 Tried manually lifting real_alpha
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 237
diff changeset
   410
  r_mk_comb_tac ctxt rty quot rel_refl trans2 rsp_thms
e9cc3a3aa5d1 Tried manually lifting real_alpha
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 237
diff changeset
   411
*}
e9cc3a3aa5d1 Tried manually lifting real_alpha
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 237
diff changeset
   412
e9cc3a3aa5d1 Tried manually lifting real_alpha
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 237
diff changeset
   413
instance lam :: fs_name
e9cc3a3aa5d1 Tried manually lifting real_alpha
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 237
diff changeset
   414
apply(intro_classes)
e9cc3a3aa5d1 Tried manually lifting real_alpha
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 237
diff changeset
   415
sorry
237
80f1df49b940 More tests in Lam
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 234
diff changeset
   416
238
e9cc3a3aa5d1 Tried manually lifting real_alpha
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 237
diff changeset
   417
prove asdf: {* Logic.mk_implies (concl_of t_r, (@{term "Trueprop (\<forall>t\<Colon>rlam\<in>Respects
e9cc3a3aa5d1 Tried manually lifting real_alpha
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 237
diff changeset
   418
          alpha.
e9cc3a3aa5d1 Tried manually lifting real_alpha
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 237
diff changeset
   419
         \<forall>(a\<Colon>name) b\<Colon>name.
e9cc3a3aa5d1 Tried manually lifting real_alpha
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 237
diff changeset
   420
            \<forall>s\<Colon>rlam\<in>Respects
e9cc3a3aa5d1 Tried manually lifting real_alpha
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 237
diff changeset
   421
                alpha.
e9cc3a3aa5d1 Tried manually lifting real_alpha
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 237
diff changeset
   422
               t \<approx> ([(a,
e9cc3a3aa5d1 Tried manually lifting real_alpha
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 237
diff changeset
   423
                 b)] \<bullet> s) \<longrightarrow>
e9cc3a3aa5d1 Tried manually lifting real_alpha
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 237
diff changeset
   424
               a = b \<or>
e9cc3a3aa5d1 Tried manually lifting real_alpha
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 237
diff changeset
   425
               a
e9cc3a3aa5d1 Tried manually lifting real_alpha
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 237
diff changeset
   426
               \<notin> {a\<Colon>name.
e9cc3a3aa5d1 Tried manually lifting real_alpha
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 237
diff changeset
   427
                infinite
e9cc3a3aa5d1 Tried manually lifting real_alpha
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 237
diff changeset
   428
                 {b\<Colon>name. Not
e9cc3a3aa5d1 Tried manually lifting real_alpha
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 237
diff changeset
   429
                  (([(a, b)] \<bullet>
e9cc3a3aa5d1 Tried manually lifting real_alpha
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 237
diff changeset
   430
                  s) \<approx>
e9cc3a3aa5d1 Tried manually lifting real_alpha
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 237
diff changeset
   431
                  s)}} \<longrightarrow>
e9cc3a3aa5d1 Tried manually lifting real_alpha
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 237
diff changeset
   432
               rLam a
e9cc3a3aa5d1 Tried manually lifting real_alpha
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 237
diff changeset
   433
                t \<approx> rLam
e9cc3a3aa5d1 Tried manually lifting real_alpha
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 237
diff changeset
   434
                 b s)"})) *}
e9cc3a3aa5d1 Tried manually lifting real_alpha
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 237
diff changeset
   435
apply (tactic {* full_simp_tac ((Simplifier.context @{context} HOL_ss) addsimps
e9cc3a3aa5d1 Tried manually lifting real_alpha
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 237
diff changeset
   436
        [(@{thm equiv_res_forall} OF [rel_eqv]),
e9cc3a3aa5d1 Tried manually lifting real_alpha
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 237
diff changeset
   437
         (@{thm equiv_res_exists} OF [rel_eqv])]) 1 *})
e9cc3a3aa5d1 Tried manually lifting real_alpha
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 237
diff changeset
   438
apply (rule allI)
e9cc3a3aa5d1 Tried manually lifting real_alpha
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 237
diff changeset
   439
apply (drule_tac x="t" in spec)
e9cc3a3aa5d1 Tried manually lifting real_alpha
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 237
diff changeset
   440
apply (rule allI)
e9cc3a3aa5d1 Tried manually lifting real_alpha
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 237
diff changeset
   441
apply (drule_tac x="a" in spec)
e9cc3a3aa5d1 Tried manually lifting real_alpha
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 237
diff changeset
   442
apply (rule allI)
e9cc3a3aa5d1 Tried manually lifting real_alpha
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 237
diff changeset
   443
apply (drule_tac x="b" in spec)
e9cc3a3aa5d1 Tried manually lifting real_alpha
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 237
diff changeset
   444
apply (rule allI)
e9cc3a3aa5d1 Tried manually lifting real_alpha
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 237
diff changeset
   445
apply (drule_tac x="s" in spec)
e9cc3a3aa5d1 Tried manually lifting real_alpha
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 237
diff changeset
   446
apply (rule impI)
e9cc3a3aa5d1 Tried manually lifting real_alpha
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 237
diff changeset
   447
apply (drule_tac mp)
e9cc3a3aa5d1 Tried manually lifting real_alpha
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 237
diff changeset
   448
apply (simp)
e9cc3a3aa5d1 Tried manually lifting real_alpha
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 237
diff changeset
   449
apply (simp)
e9cc3a3aa5d1 Tried manually lifting real_alpha
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 237
diff changeset
   450
apply (rule impI)
e9cc3a3aa5d1 Tried manually lifting real_alpha
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 237
diff changeset
   451
apply (rule a3)
e9cc3a3aa5d1 Tried manually lifting real_alpha
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 237
diff changeset
   452
apply (simp)
e9cc3a3aa5d1 Tried manually lifting real_alpha
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 237
diff changeset
   453
apply (simp add: abs_fresh(1))
e9cc3a3aa5d1 Tried manually lifting real_alpha
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 237
diff changeset
   454
apply (case_tac "a = b")
e9cc3a3aa5d1 Tried manually lifting real_alpha
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 237
diff changeset
   455
apply (simp)
e9cc3a3aa5d1 Tried manually lifting real_alpha
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 237
diff changeset
   456
apply (simp)
e9cc3a3aa5d1 Tried manually lifting real_alpha
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 237
diff changeset
   457
apply (auto)
e9cc3a3aa5d1 Tried manually lifting real_alpha
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 237
diff changeset
   458
apply (unfold fresh_def)
e9cc3a3aa5d1 Tried manually lifting real_alpha
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 237
diff changeset
   459
apply (unfold supp_def)
e9cc3a3aa5d1 Tried manually lifting real_alpha
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 237
diff changeset
   460
apply (simp)
e9cc3a3aa5d1 Tried manually lifting real_alpha
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 237
diff changeset
   461
prefer 2
e9cc3a3aa5d1 Tried manually lifting real_alpha
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 237
diff changeset
   462
apply (simp)
e9cc3a3aa5d1 Tried manually lifting real_alpha
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 237
diff changeset
   463
sorry
e9cc3a3aa5d1 Tried manually lifting real_alpha
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 237
diff changeset
   464
e9cc3a3aa5d1 Tried manually lifting real_alpha
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 237
diff changeset
   465
ML {* val abs = findabs rty (prop_of t_a) *}
e9cc3a3aa5d1 Tried manually lifting real_alpha
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 237
diff changeset
   466
ML {* val simp_lam_prs_thms = map (make_simp_lam_prs_thm @{context} quot) abs *}
e9cc3a3aa5d1 Tried manually lifting real_alpha
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 237
diff changeset
   467
ML {* val t_defs_sym = add_lower_defs @{context} defs *}
e9cc3a3aa5d1 Tried manually lifting real_alpha
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 237
diff changeset
   468
e9cc3a3aa5d1 Tried manually lifting real_alpha
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 237
diff changeset
   469
ML {* val t_r' = @{thm asdf} OF [t_r] *}
e9cc3a3aa5d1 Tried manually lifting real_alpha
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 237
diff changeset
   470
ML {* val t_t = repabs @{context} t_r' consts rty qty quot rel_refl trans2 rsp_thms *}
e9cc3a3aa5d1 Tried manually lifting real_alpha
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 237
diff changeset
   471
ML {* val t_l = repeat_eqsubst_thm @{context} simp_lam_prs_thms t_t *}
e9cc3a3aa5d1 Tried manually lifting real_alpha
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 237
diff changeset
   472
ML {* val t_a = simp_allex_prs @{context} quot t_l *}
e9cc3a3aa5d1 Tried manually lifting real_alpha
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 237
diff changeset
   473
ML {* val t_d = repeat_eqsubst_thm @{context} t_defs_sym t_a *}
e9cc3a3aa5d1 Tried manually lifting real_alpha
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 237
diff changeset
   474
ML {* val t_r = MetaSimplifier.rewrite_rule [reps_same] t_d *}
e9cc3a3aa5d1 Tried manually lifting real_alpha
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 237
diff changeset
   475
ML {* val tt = MetaSimplifier.rewrite_rule [symmetric @{thm supp_def}, symmetric @{thm fresh_def}] t_r *}
e9cc3a3aa5d1 Tried manually lifting real_alpha
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 237
diff changeset
   476
ML {* val rr = @{thm sym} OF @{thms abs_fresh(1)} *}
e9cc3a3aa5d1 Tried manually lifting real_alpha
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 237
diff changeset
   477
ML {* val ttt = eqsubst_thm @{context} [rr] tt *}
e9cc3a3aa5d1 Tried manually lifting real_alpha
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 237
diff changeset
   478
ML {* ObjectLogic.rulify ttt *}
e9cc3a3aa5d1 Tried manually lifting real_alpha
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 237
diff changeset
   479
e9cc3a3aa5d1 Tried manually lifting real_alpha
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 237
diff changeset
   480
lemma
e9cc3a3aa5d1 Tried manually lifting real_alpha
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 237
diff changeset
   481
  assumes a: "a \<notin> {a\<Colon>name. infinite {b\<Colon>name. \<not> ([(a, b)] \<bullet> s) \<approx> s}}"
e9cc3a3aa5d1 Tried manually lifting real_alpha
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 237
diff changeset
   482
  shows "a \<notin> {a\<Colon>name. infinite {b\<Colon>name. [(a, b)] \<bullet> s \<noteq> s}}"
e9cc3a3aa5d1 Tried manually lifting real_alpha
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 237
diff changeset
   483
  using a apply simp
e9cc3a3aa5d1 Tried manually lifting real_alpha
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 237
diff changeset
   484
  sorry (* Not true... *)