Quot/Examples/LamEx.thy
author Christian Urban <urbanc@in.tum.de>
Fri, 15 Jan 2010 17:09:36 +0100
changeset 895 92c43c96027e
parent 894 1d80641a4302
child 896 4e0b89d886ac
permissions -rw-r--r--
added a partial proof under which conditions rlam_rec Respects alpha...I guess something like this is true; this means the Hom lemmas need to have preconditions
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
758
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents: 705
diff changeset
     2
imports Nominal "../QuotMain" "../QuotList"
201
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
804
Christian Urban <urbanc@in.tum.de>
parents: 767
diff changeset
     7
datatype rlam =
201
1ac36993cc71 added an example about lambda-terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     8
  rVar "name"
1ac36993cc71 added an example about lambda-terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     9
| rApp "rlam" "rlam"
1ac36993cc71 added an example about lambda-terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    10
| rLam "name" "rlam"
1ac36993cc71 added an example about lambda-terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    11
804
Christian Urban <urbanc@in.tum.de>
parents: 767
diff changeset
    12
fun
243
22715cab3995 added fv-function
Christian Urban <urbanc@in.tum.de>
parents: 240
diff changeset
    13
  rfv :: "rlam \<Rightarrow> name set"
22715cab3995 added fv-function
Christian Urban <urbanc@in.tum.de>
parents: 240
diff changeset
    14
where
22715cab3995 added fv-function
Christian Urban <urbanc@in.tum.de>
parents: 240
diff changeset
    15
  rfv_var: "rfv (rVar a) = {a}"
22715cab3995 added fv-function
Christian Urban <urbanc@in.tum.de>
parents: 240
diff changeset
    16
| 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
    17
| rfv_lam: "rfv (rLam a t) = (rfv t) - {a}"
804
Christian Urban <urbanc@in.tum.de>
parents: 767
diff changeset
    18
Christian Urban <urbanc@in.tum.de>
parents: 767
diff changeset
    19
overloading
876
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 804
diff changeset
    20
  perm_rlam \<equiv> "perm :: 'x prm \<Rightarrow> rlam \<Rightarrow> rlam"   (unchecked)
804
Christian Urban <urbanc@in.tum.de>
parents: 767
diff changeset
    21
begin
243
22715cab3995 added fv-function
Christian Urban <urbanc@in.tum.de>
parents: 240
diff changeset
    22
804
Christian Urban <urbanc@in.tum.de>
parents: 767
diff changeset
    23
fun
Christian Urban <urbanc@in.tum.de>
parents: 767
diff changeset
    24
  perm_rlam
Christian Urban <urbanc@in.tum.de>
parents: 767
diff changeset
    25
where
Christian Urban <urbanc@in.tum.de>
parents: 767
diff changeset
    26
  "perm_rlam pi (rVar a) = rVar (pi \<bullet> a)"
Christian Urban <urbanc@in.tum.de>
parents: 767
diff changeset
    27
| "perm_rlam pi (rApp t1 t2) = rApp (perm_rlam pi t1) (perm_rlam pi t2)"
Christian Urban <urbanc@in.tum.de>
parents: 767
diff changeset
    28
| "perm_rlam pi (rLam a t) = rLam (pi \<bullet> a) (perm_rlam pi t)"
Christian Urban <urbanc@in.tum.de>
parents: 767
diff changeset
    29
Christian Urban <urbanc@in.tum.de>
parents: 767
diff changeset
    30
end
243
22715cab3995 added fv-function
Christian Urban <urbanc@in.tum.de>
parents: 240
diff changeset
    31
895
92c43c96027e added a partial proof under which conditions rlam_rec Respects alpha...I guess something like this is true; this means the Hom lemmas need to have preconditions
Christian Urban <urbanc@in.tum.de>
parents: 894
diff changeset
    32
instance rlam::pt_name
92c43c96027e added a partial proof under which conditions rlam_rec Respects alpha...I guess something like this is true; this means the Hom lemmas need to have preconditions
Christian Urban <urbanc@in.tum.de>
parents: 894
diff changeset
    33
apply(default)
92c43c96027e added a partial proof under which conditions rlam_rec Respects alpha...I guess something like this is true; this means the Hom lemmas need to have preconditions
Christian Urban <urbanc@in.tum.de>
parents: 894
diff changeset
    34
apply(induct_tac [!] x rule: rlam.induct)
92c43c96027e added a partial proof under which conditions rlam_rec Respects alpha...I guess something like this is true; this means the Hom lemmas need to have preconditions
Christian Urban <urbanc@in.tum.de>
parents: 894
diff changeset
    35
apply(simp_all add: pt_name2 pt_name3)
92c43c96027e added a partial proof under which conditions rlam_rec Respects alpha...I guess something like this is true; this means the Hom lemmas need to have preconditions
Christian Urban <urbanc@in.tum.de>
parents: 894
diff changeset
    36
done
92c43c96027e added a partial proof under which conditions rlam_rec Respects alpha...I guess something like this is true; this means the Hom lemmas need to have preconditions
Christian Urban <urbanc@in.tum.de>
parents: 894
diff changeset
    37
92c43c96027e added a partial proof under which conditions rlam_rec Respects alpha...I guess something like this is true; this means the Hom lemmas need to have preconditions
Christian Urban <urbanc@in.tum.de>
parents: 894
diff changeset
    38
instance rlam::fs_name
92c43c96027e added a partial proof under which conditions rlam_rec Respects alpha...I guess something like this is true; this means the Hom lemmas need to have preconditions
Christian Urban <urbanc@in.tum.de>
parents: 894
diff changeset
    39
apply(default)
92c43c96027e added a partial proof under which conditions rlam_rec Respects alpha...I guess something like this is true; this means the Hom lemmas need to have preconditions
Christian Urban <urbanc@in.tum.de>
parents: 894
diff changeset
    40
apply(induct_tac [!] x rule: rlam.induct)
92c43c96027e added a partial proof under which conditions rlam_rec Respects alpha...I guess something like this is true; this means the Hom lemmas need to have preconditions
Christian Urban <urbanc@in.tum.de>
parents: 894
diff changeset
    41
apply(simp add: supp_def)
92c43c96027e added a partial proof under which conditions rlam_rec Respects alpha...I guess something like this is true; this means the Hom lemmas need to have preconditions
Christian Urban <urbanc@in.tum.de>
parents: 894
diff changeset
    42
apply(fold supp_def)
92c43c96027e added a partial proof under which conditions rlam_rec Respects alpha...I guess something like this is true; this means the Hom lemmas need to have preconditions
Christian Urban <urbanc@in.tum.de>
parents: 894
diff changeset
    43
apply(simp add: supp_atm)
92c43c96027e added a partial proof under which conditions rlam_rec Respects alpha...I guess something like this is true; this means the Hom lemmas need to have preconditions
Christian Urban <urbanc@in.tum.de>
parents: 894
diff changeset
    44
apply(simp add: supp_def Collect_imp_eq Collect_neg_eq)
92c43c96027e added a partial proof under which conditions rlam_rec Respects alpha...I guess something like this is true; this means the Hom lemmas need to have preconditions
Christian Urban <urbanc@in.tum.de>
parents: 894
diff changeset
    45
apply(simp add: supp_def)
92c43c96027e added a partial proof under which conditions rlam_rec Respects alpha...I guess something like this is true; this means the Hom lemmas need to have preconditions
Christian Urban <urbanc@in.tum.de>
parents: 894
diff changeset
    46
apply(simp add: supp_def Collect_imp_eq Collect_neg_eq[symmetric])
92c43c96027e added a partial proof under which conditions rlam_rec Respects alpha...I guess something like this is true; this means the Hom lemmas need to have preconditions
Christian Urban <urbanc@in.tum.de>
parents: 894
diff changeset
    47
apply(fold supp_def)
92c43c96027e added a partial proof under which conditions rlam_rec Respects alpha...I guess something like this is true; this means the Hom lemmas need to have preconditions
Christian Urban <urbanc@in.tum.de>
parents: 894
diff changeset
    48
apply(simp add: supp_atm)
92c43c96027e added a partial proof under which conditions rlam_rec Respects alpha...I guess something like this is true; this means the Hom lemmas need to have preconditions
Christian Urban <urbanc@in.tum.de>
parents: 894
diff changeset
    49
done
92c43c96027e added a partial proof under which conditions rlam_rec Respects alpha...I guess something like this is true; this means the Hom lemmas need to have preconditions
Christian Urban <urbanc@in.tum.de>
parents: 894
diff changeset
    50
271
1b57f99737fe Alpha.induct now lifts automatically.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 270
diff changeset
    51
inductive
246
6da7d538e5e0 changed the order of rfv and reformulated a3 with rfv
Christian Urban <urbanc@in.tum.de>
parents: 245
diff changeset
    52
  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
    53
where
6da7d538e5e0 changed the order of rfv and reformulated a3 with rfv
Christian Urban <urbanc@in.tum.de>
parents: 245
diff changeset
    54
  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
    55
| a2: "\<lbrakk>t1 \<approx> t2; s1 \<approx> s2\<rbrakk> \<Longrightarrow> rApp t1 s1 \<approx> rApp t2 s2"
804
Christian Urban <urbanc@in.tum.de>
parents: 767
diff changeset
    56
| a3: "\<lbrakk>t \<approx> ([(a,b)] \<bullet> s); a \<notin> rfv (rLam b t)\<rbrakk> \<Longrightarrow> rLam a t \<approx> rLam b s"
246
6da7d538e5e0 changed the order of rfv and reformulated a3 with rfv
Christian Urban <urbanc@in.tum.de>
parents: 245
diff changeset
    57
804
Christian Urban <urbanc@in.tum.de>
parents: 767
diff changeset
    58
lemma helper:
Christian Urban <urbanc@in.tum.de>
parents: 767
diff changeset
    59
  fixes t::"rlam"
Christian Urban <urbanc@in.tum.de>
parents: 767
diff changeset
    60
  and   a::"name"
Christian Urban <urbanc@in.tum.de>
parents: 767
diff changeset
    61
  shows "[(a, a)] \<bullet> t = t"
Christian Urban <urbanc@in.tum.de>
parents: 767
diff changeset
    62
by (induct t)
Christian Urban <urbanc@in.tum.de>
parents: 767
diff changeset
    63
   (auto simp add: calc_atm)
259
22c199522bef Optimization
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 258
diff changeset
    64
271
1b57f99737fe Alpha.induct now lifts automatically.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 270
diff changeset
    65
lemma alpha_refl:
286
Christian Urban <urbanc@in.tum.de>
parents: 285
diff changeset
    66
  fixes t::"rlam"
Christian Urban <urbanc@in.tum.de>
parents: 285
diff changeset
    67
  shows "t \<approx> t"
Christian Urban <urbanc@in.tum.de>
parents: 285
diff changeset
    68
  apply(induct t rule: rlam.induct)
Christian Urban <urbanc@in.tum.de>
parents: 285
diff changeset
    69
  apply(simp add: a1)
Christian Urban <urbanc@in.tum.de>
parents: 285
diff changeset
    70
  apply(simp add: a2)
Christian Urban <urbanc@in.tum.de>
parents: 285
diff changeset
    71
  apply(rule a3)
804
Christian Urban <urbanc@in.tum.de>
parents: 767
diff changeset
    72
  apply(simp add: helper)
286
Christian Urban <urbanc@in.tum.de>
parents: 285
diff changeset
    73
  apply(simp)
Christian Urban <urbanc@in.tum.de>
parents: 285
diff changeset
    74
  done
Christian Urban <urbanc@in.tum.de>
parents: 285
diff changeset
    75
534
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 515
diff changeset
    76
lemma alpha_equivp:
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 515
diff changeset
    77
  shows "equivp alpha"
271
1b57f99737fe Alpha.induct now lifts automatically.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 270
diff changeset
    78
sorry
1b57f99737fe Alpha.induct now lifts automatically.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 270
diff changeset
    79
766
df053507edba renamed "quotient" command to "quotient_type"; needs new keyword file to be installed
Christian Urban <urbanc@in.tum.de>
parents: 758
diff changeset
    80
quotient_type lam = rlam / alpha
804
Christian Urban <urbanc@in.tum.de>
parents: 767
diff changeset
    81
  by (rule alpha_equivp)
201
1ac36993cc71 added an example about lambda-terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    82
203
7384115df9fd added equiv-thm to the quot_info
Christian Urban <urbanc@in.tum.de>
parents: 201
diff changeset
    83
767
37285ec4387d on suggestion of Tobias renamed "quotient_def" to "quotient_definition"; needs new keyword file
Christian Urban <urbanc@in.tum.de>
parents: 766
diff changeset
    84
quotient_definition
876
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 804
diff changeset
    85
  "Var :: name \<Rightarrow> lam"
705
f51c6069cd17 New syntax for definitions.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 663
diff changeset
    86
as
663
0dd10a900cae Different syntax for definitions that allows overloading and retrieving of definitions by matching whole constants.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 636
diff changeset
    87
  "rVar"
229
13f985a93dbc fixed the definition of alpha; this *breaks* some of the experiments
Christian Urban <urbanc@in.tum.de>
parents: 225
diff changeset
    88
767
37285ec4387d on suggestion of Tobias renamed "quotient_def" to "quotient_definition"; needs new keyword file
Christian Urban <urbanc@in.tum.de>
parents: 766
diff changeset
    89
quotient_definition
705
f51c6069cd17 New syntax for definitions.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 663
diff changeset
    90
   "App :: lam \<Rightarrow> lam \<Rightarrow> lam"
f51c6069cd17 New syntax for definitions.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 663
diff changeset
    91
as
663
0dd10a900cae Different syntax for definitions that allows overloading and retrieving of definitions by matching whole constants.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 636
diff changeset
    92
  "rApp"
229
13f985a93dbc fixed the definition of alpha; this *breaks* some of the experiments
Christian Urban <urbanc@in.tum.de>
parents: 225
diff changeset
    93
767
37285ec4387d on suggestion of Tobias renamed "quotient_def" to "quotient_definition"; needs new keyword file
Christian Urban <urbanc@in.tum.de>
parents: 766
diff changeset
    94
quotient_definition
876
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 804
diff changeset
    95
  "Lam :: name \<Rightarrow> lam \<Rightarrow> lam"
705
f51c6069cd17 New syntax for definitions.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 663
diff changeset
    96
as
663
0dd10a900cae Different syntax for definitions that allows overloading and retrieving of definitions by matching whole constants.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 636
diff changeset
    97
  "rLam"
201
1ac36993cc71 added an example about lambda-terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    98
767
37285ec4387d on suggestion of Tobias renamed "quotient_def" to "quotient_definition"; needs new keyword file
Christian Urban <urbanc@in.tum.de>
parents: 766
diff changeset
    99
quotient_definition
876
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 804
diff changeset
   100
  "fv :: lam \<Rightarrow> name set"
705
f51c6069cd17 New syntax for definitions.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 663
diff changeset
   101
as
663
0dd10a900cae Different syntax for definitions that allows overloading and retrieving of definitions by matching whole constants.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 636
diff changeset
   102
  "rfv"
243
22715cab3995 added fv-function
Christian Urban <urbanc@in.tum.de>
parents: 240
diff changeset
   103
229
13f985a93dbc fixed the definition of alpha; this *breaks* some of the experiments
Christian Urban <urbanc@in.tum.de>
parents: 225
diff changeset
   104
(* 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
   105
(* 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
   106
overloading
268
4d58c02289ca simplified the quotient_def code; type of the defined constant must now be given; for-part eliminated
Christian Urban <urbanc@in.tum.de>
parents: 267
diff changeset
   107
  perm_lam \<equiv> "perm :: 'x prm \<Rightarrow> lam \<Rightarrow> lam"   (unchecked)
229
13f985a93dbc fixed the definition of alpha; this *breaks* some of the experiments
Christian Urban <urbanc@in.tum.de>
parents: 225
diff changeset
   108
begin
13f985a93dbc fixed the definition of alpha; this *breaks* some of the experiments
Christian Urban <urbanc@in.tum.de>
parents: 225
diff changeset
   109
767
37285ec4387d on suggestion of Tobias renamed "quotient_def" to "quotient_definition"; needs new keyword file
Christian Urban <urbanc@in.tum.de>
parents: 766
diff changeset
   110
quotient_definition
705
f51c6069cd17 New syntax for definitions.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 663
diff changeset
   111
   "perm_lam :: 'x prm \<Rightarrow> lam \<Rightarrow> lam"
f51c6069cd17 New syntax for definitions.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 663
diff changeset
   112
as
663
0dd10a900cae Different syntax for definitions that allows overloading and retrieving of definitions by matching whole constants.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 636
diff changeset
   113
  "perm::'x prm \<Rightarrow> rlam \<Rightarrow> rlam"
229
13f985a93dbc fixed the definition of alpha; this *breaks* some of the experiments
Christian Urban <urbanc@in.tum.de>
parents: 225
diff changeset
   114
13f985a93dbc fixed the definition of alpha; this *breaks* some of the experiments
Christian Urban <urbanc@in.tum.de>
parents: 225
diff changeset
   115
end
13f985a93dbc fixed the definition of alpha; this *breaks* some of the experiments
Christian Urban <urbanc@in.tum.de>
parents: 225
diff changeset
   116
804
Christian Urban <urbanc@in.tum.de>
parents: 767
diff changeset
   117
(* lemmas that need to be lifted *)
Christian Urban <urbanc@in.tum.de>
parents: 767
diff changeset
   118
lemma pi_var_eqvt1:
229
13f985a93dbc fixed the definition of alpha; this *breaks* some of the experiments
Christian Urban <urbanc@in.tum.de>
parents: 225
diff changeset
   119
  fixes pi::"'x prm"
804
Christian Urban <urbanc@in.tum.de>
parents: 767
diff changeset
   120
  shows "(pi \<bullet> rVar a) \<approx> rVar (pi \<bullet> a)"
Christian Urban <urbanc@in.tum.de>
parents: 767
diff changeset
   121
  by (simp add: alpha_refl)
229
13f985a93dbc fixed the definition of alpha; this *breaks* some of the experiments
Christian Urban <urbanc@in.tum.de>
parents: 225
diff changeset
   122
804
Christian Urban <urbanc@in.tum.de>
parents: 767
diff changeset
   123
lemma pi_var_eqvt2:
Christian Urban <urbanc@in.tum.de>
parents: 767
diff changeset
   124
  fixes pi::"'x prm"
Christian Urban <urbanc@in.tum.de>
parents: 767
diff changeset
   125
  shows "(pi \<bullet> rVar a) = rVar (pi \<bullet> a)"
Christian Urban <urbanc@in.tum.de>
parents: 767
diff changeset
   126
  by (simp)
Christian Urban <urbanc@in.tum.de>
parents: 767
diff changeset
   127
Christian Urban <urbanc@in.tum.de>
parents: 767
diff changeset
   128
lemma pi_app_eqvt1:
229
13f985a93dbc fixed the definition of alpha; this *breaks* some of the experiments
Christian Urban <urbanc@in.tum.de>
parents: 225
diff changeset
   129
  fixes pi::"'x prm"
804
Christian Urban <urbanc@in.tum.de>
parents: 767
diff changeset
   130
  shows "(pi \<bullet> rApp t1 t2) \<approx> rApp (pi \<bullet> t1) (pi \<bullet> t2)"
Christian Urban <urbanc@in.tum.de>
parents: 767
diff changeset
   131
  by (simp add: alpha_refl)
Christian Urban <urbanc@in.tum.de>
parents: 767
diff changeset
   132
Christian Urban <urbanc@in.tum.de>
parents: 767
diff changeset
   133
lemma pi_app_eqvt2:
Christian Urban <urbanc@in.tum.de>
parents: 767
diff changeset
   134
  fixes pi::"'x prm"
Christian Urban <urbanc@in.tum.de>
parents: 767
diff changeset
   135
  shows "(pi \<bullet> rApp t1 t2) = rApp (pi \<bullet> t1) (pi \<bullet> t2)"
Christian Urban <urbanc@in.tum.de>
parents: 767
diff changeset
   136
  by (simp)
229
13f985a93dbc fixed the definition of alpha; this *breaks* some of the experiments
Christian Urban <urbanc@in.tum.de>
parents: 225
diff changeset
   137
804
Christian Urban <urbanc@in.tum.de>
parents: 767
diff changeset
   138
lemma pi_lam_eqvt1:
229
13f985a93dbc fixed the definition of alpha; this *breaks* some of the experiments
Christian Urban <urbanc@in.tum.de>
parents: 225
diff changeset
   139
  fixes pi::"'x prm"
804
Christian Urban <urbanc@in.tum.de>
parents: 767
diff changeset
   140
  shows "(pi \<bullet> rLam a t) \<approx> rLam (pi \<bullet> a) (pi \<bullet> t)"
Christian Urban <urbanc@in.tum.de>
parents: 767
diff changeset
   141
  by (simp add: alpha_refl)
229
13f985a93dbc fixed the definition of alpha; this *breaks* some of the experiments
Christian Urban <urbanc@in.tum.de>
parents: 225
diff changeset
   142
804
Christian Urban <urbanc@in.tum.de>
parents: 767
diff changeset
   143
lemma pi_lam_eqvt2:
Christian Urban <urbanc@in.tum.de>
parents: 767
diff changeset
   144
  fixes pi::"'x prm"
Christian Urban <urbanc@in.tum.de>
parents: 767
diff changeset
   145
  shows "(pi \<bullet> rLam a t) = rLam (pi \<bullet> a) (pi \<bullet> t)"
Christian Urban <urbanc@in.tum.de>
parents: 767
diff changeset
   146
  by (simp add: alpha)
451
586e3dc4afdb Added 'TRY' to refl in clean_tac to get as far as possible. Removed unnecessary [quot_rsp] in FSet. Added necessary [quot_rsp] and one lifted thm in LamEx.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 445
diff changeset
   147
586e3dc4afdb Added 'TRY' to refl in clean_tac to get as far as possible. Removed unnecessary [quot_rsp] in FSet. Added necessary [quot_rsp] and one lifted thm in LamEx.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 445
diff changeset
   148
201
1ac36993cc71 added an example about lambda-terms
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   149
lemma real_alpha:
451
586e3dc4afdb Added 'TRY' to refl in clean_tac to get as far as possible. Removed unnecessary [quot_rsp] in FSet. Added necessary [quot_rsp] and one lifted thm in LamEx.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 445
diff changeset
   150
  assumes a: "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
   151
  shows "Lam a t = Lam b s"
451
586e3dc4afdb Added 'TRY' to refl in clean_tac to get as far as possible. Removed unnecessary [quot_rsp] in FSet. Added necessary [quot_rsp] and one lifted thm in LamEx.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 445
diff changeset
   152
using a
586e3dc4afdb Added 'TRY' to refl in clean_tac to get as far as possible. Removed unnecessary [quot_rsp] in FSet. Added necessary [quot_rsp] and one lifted thm in LamEx.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 445
diff changeset
   153
unfolding fresh_def supp_def
217
9cc87d02190a First experiments with Lambda
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 203
diff changeset
   154
sorry
9cc87d02190a First experiments with Lambda
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 203
diff changeset
   155
636
520a4084d064 changed names of attributes
Christian Urban <urbanc@in.tum.de>
parents: 610
diff changeset
   156
lemma perm_rsp[quot_respect]:
286
Christian Urban <urbanc@in.tum.de>
parents: 285
diff changeset
   157
  "(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
   158
  apply(auto)
13f985a93dbc fixed the definition of alpha; this *breaks* some of the experiments
Christian Urban <urbanc@in.tum.de>
parents: 225
diff changeset
   159
  (* 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
   160
  sorry
13f985a93dbc fixed the definition of alpha; this *breaks* some of the experiments
Christian Urban <urbanc@in.tum.de>
parents: 225
diff changeset
   161
451
586e3dc4afdb Added 'TRY' to refl in clean_tac to get as far as possible. Removed unnecessary [quot_rsp] in FSet. Added necessary [quot_rsp] and one lifted thm in LamEx.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 445
diff changeset
   162
lemma fresh_rsp:
586e3dc4afdb Added 'TRY' to refl in clean_tac to get as far as possible. Removed unnecessary [quot_rsp] in FSet. Added necessary [quot_rsp] and one lifted thm in LamEx.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 445
diff changeset
   163
  "(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
   164
  apply(auto)
13f985a93dbc fixed the definition of alpha; this *breaks* some of the experiments
Christian Urban <urbanc@in.tum.de>
parents: 225
diff changeset
   165
  (* 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
   166
  sorry
13f985a93dbc fixed the definition of alpha; this *breaks* some of the experiments
Christian Urban <urbanc@in.tum.de>
parents: 225
diff changeset
   167
636
520a4084d064 changed names of attributes
Christian Urban <urbanc@in.tum.de>
parents: 610
diff changeset
   168
lemma rVar_rsp[quot_respect]:
451
586e3dc4afdb Added 'TRY' to refl in clean_tac to get as far as possible. Removed unnecessary [quot_rsp] in FSet. Added necessary [quot_rsp] and one lifted thm in LamEx.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 445
diff changeset
   169
  "(op = ===> alpha) rVar rVar"
876
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 804
diff changeset
   170
  by (auto intro: a1)
234
811f17de4031 Lifting of the 3 lemmas in LamEx
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 233
diff changeset
   171
636
520a4084d064 changed names of attributes
Christian Urban <urbanc@in.tum.de>
parents: 610
diff changeset
   172
lemma rApp_rsp[quot_respect]: "(alpha ===> alpha ===> alpha) rApp rApp"
876
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 804
diff changeset
   173
  by (auto intro: a2)
234
811f17de4031 Lifting of the 3 lemmas in LamEx
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 233
diff changeset
   174
636
520a4084d064 changed names of attributes
Christian Urban <urbanc@in.tum.de>
parents: 610
diff changeset
   175
lemma rLam_rsp[quot_respect]: "(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
   176
  apply(auto)
13f985a93dbc fixed the definition of alpha; this *breaks* some of the experiments
Christian Urban <urbanc@in.tum.de>
parents: 225
diff changeset
   177
  apply(rule a3)
804
Christian Urban <urbanc@in.tum.de>
parents: 767
diff changeset
   178
  apply(simp add: helper)
Christian Urban <urbanc@in.tum.de>
parents: 767
diff changeset
   179
  apply(simp)
229
13f985a93dbc fixed the definition of alpha; this *breaks* some of the experiments
Christian Urban <urbanc@in.tum.de>
parents: 225
diff changeset
   180
  done
217
9cc87d02190a First experiments with Lambda
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 203
diff changeset
   181
804
Christian Urban <urbanc@in.tum.de>
parents: 767
diff changeset
   182
lemma rfv_rsp[quot_respect]: 
Christian Urban <urbanc@in.tum.de>
parents: 767
diff changeset
   183
  "(alpha ===> op =) rfv rfv"
247
e83a6e452843 Lemmas about fv.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 246
diff changeset
   184
  sorry
217
9cc87d02190a First experiments with Lambda
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 203
diff changeset
   185
285
8ebdef196fd5 Infrastructure for polymorphic types
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 273
diff changeset
   186
lemma rvar_inject: "rVar a \<approx> rVar b = (a = b)"
876
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 804
diff changeset
   187
  apply (auto)
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 804
diff changeset
   188
  apply (erule alpha.cases)
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 804
diff changeset
   189
  apply (simp_all add: rlam.inject alpha_refl)
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 804
diff changeset
   190
  done
237
80f1df49b940 More tests in Lam
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 234
diff changeset
   191
804
Christian Urban <urbanc@in.tum.de>
parents: 767
diff changeset
   192
lemma pi_var1:
Christian Urban <urbanc@in.tum.de>
parents: 767
diff changeset
   193
  fixes pi::"'x prm"
Christian Urban <urbanc@in.tum.de>
parents: 767
diff changeset
   194
  shows "pi \<bullet> Var a = Var (pi \<bullet> a)"
Christian Urban <urbanc@in.tum.de>
parents: 767
diff changeset
   195
  by (lifting pi_var_eqvt1)
370
09e28d4c19aa Lambda & SOLVED' for new quotient_tac
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 292
diff changeset
   196
804
Christian Urban <urbanc@in.tum.de>
parents: 767
diff changeset
   197
lemma pi_var2:
Christian Urban <urbanc@in.tum.de>
parents: 767
diff changeset
   198
  fixes pi::"'x prm"
Christian Urban <urbanc@in.tum.de>
parents: 767
diff changeset
   199
  shows "pi \<bullet> Var a = Var (pi \<bullet> a)"
Christian Urban <urbanc@in.tum.de>
parents: 767
diff changeset
   200
  by (lifting pi_var_eqvt2)
Christian Urban <urbanc@in.tum.de>
parents: 767
diff changeset
   201
376
e99c0334d8bf lambda_prs and cleaning the existing examples.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 370
diff changeset
   202
804
Christian Urban <urbanc@in.tum.de>
parents: 767
diff changeset
   203
lemma pi_app: 
Christian Urban <urbanc@in.tum.de>
parents: 767
diff changeset
   204
  fixes pi::"'x prm"
Christian Urban <urbanc@in.tum.de>
parents: 767
diff changeset
   205
  shows "pi \<bullet> App t1 t2 = App (pi \<bullet> t1) (pi \<bullet> t2)"
Christian Urban <urbanc@in.tum.de>
parents: 767
diff changeset
   206
  by (lifting pi_app_eqvt2)
249
7dec34d12328 added some facts about fresh and support of lam
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
   207
804
Christian Urban <urbanc@in.tum.de>
parents: 767
diff changeset
   208
lemma pi_lam: 
Christian Urban <urbanc@in.tum.de>
parents: 767
diff changeset
   209
  fixes pi::"'x prm"
Christian Urban <urbanc@in.tum.de>
parents: 767
diff changeset
   210
  shows "pi \<bullet> Lam a t = Lam (pi \<bullet> a) (pi \<bullet> t)"
Christian Urban <urbanc@in.tum.de>
parents: 767
diff changeset
   211
  by (lifting pi_lam_eqvt2)
376
e99c0334d8bf lambda_prs and cleaning the existing examples.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 370
diff changeset
   212
804
Christian Urban <urbanc@in.tum.de>
parents: 767
diff changeset
   213
lemma fv_var: 
Christian Urban <urbanc@in.tum.de>
parents: 767
diff changeset
   214
  shows "fv (Var a) = {a}"
Christian Urban <urbanc@in.tum.de>
parents: 767
diff changeset
   215
  by  (lifting rfv_var)
249
7dec34d12328 added some facts about fresh and support of lam
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
   216
804
Christian Urban <urbanc@in.tum.de>
parents: 767
diff changeset
   217
lemma fv_app: 
Christian Urban <urbanc@in.tum.de>
parents: 767
diff changeset
   218
  shows "fv (App t1 t2) = fv t1 \<union> fv t2"
Christian Urban <urbanc@in.tum.de>
parents: 767
diff changeset
   219
  by (lifting rfv_app)
Christian Urban <urbanc@in.tum.de>
parents: 767
diff changeset
   220
Christian Urban <urbanc@in.tum.de>
parents: 767
diff changeset
   221
lemma fv_lam: 
Christian Urban <urbanc@in.tum.de>
parents: 767
diff changeset
   222
  shows "fv (Lam a t) = fv t - {a}"
Christian Urban <urbanc@in.tum.de>
parents: 767
diff changeset
   223
  by (lifting rfv_lam)
376
e99c0334d8bf lambda_prs and cleaning the existing examples.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 370
diff changeset
   224
804
Christian Urban <urbanc@in.tum.de>
parents: 767
diff changeset
   225
lemma a1: 
Christian Urban <urbanc@in.tum.de>
parents: 767
diff changeset
   226
  "a = b \<Longrightarrow> Var a = Var b"
Christian Urban <urbanc@in.tum.de>
parents: 767
diff changeset
   227
  by  (lifting a1)
376
e99c0334d8bf lambda_prs and cleaning the existing examples.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 370
diff changeset
   228
804
Christian Urban <urbanc@in.tum.de>
parents: 767
diff changeset
   229
lemma a2: 
Christian Urban <urbanc@in.tum.de>
parents: 767
diff changeset
   230
  "\<lbrakk>x = xa; xb = xc\<rbrakk> \<Longrightarrow> App x xb = App xa xc"
Christian Urban <urbanc@in.tum.de>
parents: 767
diff changeset
   231
  by  (lifting a2)
376
e99c0334d8bf lambda_prs and cleaning the existing examples.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 370
diff changeset
   232
804
Christian Urban <urbanc@in.tum.de>
parents: 767
diff changeset
   233
lemma a3: 
Christian Urban <urbanc@in.tum.de>
parents: 767
diff changeset
   234
  "\<lbrakk>x = [(a, b)] \<bullet> xa; a \<notin> fv (Lam b x)\<rbrakk> \<Longrightarrow> Lam a x = Lam b xa"
Christian Urban <urbanc@in.tum.de>
parents: 767
diff changeset
   235
  by  (lifting a3)
286
Christian Urban <urbanc@in.tum.de>
parents: 285
diff changeset
   236
804
Christian Urban <urbanc@in.tum.de>
parents: 767
diff changeset
   237
lemma alpha_cases: 
Christian Urban <urbanc@in.tum.de>
parents: 767
diff changeset
   238
  "\<lbrakk>a1 = a2; \<And>a b. \<lbrakk>a1 = Var a; a2 = Var b; a = b\<rbrakk> \<Longrightarrow> P;
Christian Urban <urbanc@in.tum.de>
parents: 767
diff changeset
   239
    \<And>x xa xb xc. \<lbrakk>a1 = App x xb; a2 = App xa xc; x = xa; xb = xc\<rbrakk> \<Longrightarrow> P;
Christian Urban <urbanc@in.tum.de>
parents: 767
diff changeset
   240
    \<And>x a b xa. \<lbrakk>a1 = Lam a x; a2 = Lam b xa; x = [(a, b)] \<bullet> xa; a \<notin> fv (Lam b x)\<rbrakk> \<Longrightarrow> P\<rbrakk>
370
09e28d4c19aa Lambda & SOLVED' for new quotient_tac
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 292
diff changeset
   241
    \<Longrightarrow> P"
804
Christian Urban <urbanc@in.tum.de>
parents: 767
diff changeset
   242
  by (lifting alpha.cases)
378
86fba2c4eeef All examples work again.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 376
diff changeset
   243
804
Christian Urban <urbanc@in.tum.de>
parents: 767
diff changeset
   244
lemma alpha_induct: 
Christian Urban <urbanc@in.tum.de>
parents: 767
diff changeset
   245
  "\<lbrakk>qx = qxa; \<And>a b. a = b \<Longrightarrow> qxb (Var a) (Var b);
Christian Urban <urbanc@in.tum.de>
parents: 767
diff changeset
   246
    \<And>x xa xb xc. \<lbrakk>x = xa; qxb x xa; xb = xc; qxb xb xc\<rbrakk> \<Longrightarrow> qxb (App x xb) (App xa xc);
Christian Urban <urbanc@in.tum.de>
parents: 767
diff changeset
   247
     \<And>x a b xa.
370
09e28d4c19aa Lambda & SOLVED' for new quotient_tac
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 292
diff changeset
   248
        \<lbrakk>x = [(a, b)] \<bullet> xa; qxb x ([(a, b)] \<bullet> xa); a \<notin> fv (Lam b x)\<rbrakk> \<Longrightarrow> qxb (Lam a x) (Lam b xa)\<rbrakk>
09e28d4c19aa Lambda & SOLVED' for new quotient_tac
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 292
diff changeset
   249
    \<Longrightarrow> qxb qx qxa"
804
Christian Urban <urbanc@in.tum.de>
parents: 767
diff changeset
   250
  by (lifting alpha.induct)
Christian Urban <urbanc@in.tum.de>
parents: 767
diff changeset
   251
Christian Urban <urbanc@in.tum.de>
parents: 767
diff changeset
   252
lemma var_inject: 
Christian Urban <urbanc@in.tum.de>
parents: 767
diff changeset
   253
  "(Var a = Var b) = (a = b)"
Christian Urban <urbanc@in.tum.de>
parents: 767
diff changeset
   254
  by (lifting rvar_inject)
252
e30997c88050 Regularize for equalities and a better tactic. "alpha.cases" now lifts.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 251
diff changeset
   255
883
99e811fc1366 a few more lemmas...except supp of lambda-abstractions
Christian Urban <urbanc@in.tum.de>
parents: 882
diff changeset
   256
99e811fc1366 a few more lemmas...except supp of lambda-abstractions
Christian Urban <urbanc@in.tum.de>
parents: 882
diff changeset
   257
lemma app_inject: 
99e811fc1366 a few more lemmas...except supp of lambda-abstractions
Christian Urban <urbanc@in.tum.de>
parents: 882
diff changeset
   258
  "(App t1 t2 = App s1 s2) = (t1 = s1 \<and> t2 = s2)"
99e811fc1366 a few more lemmas...except supp of lambda-abstractions
Christian Urban <urbanc@in.tum.de>
parents: 882
diff changeset
   259
sorry
99e811fc1366 a few more lemmas...except supp of lambda-abstractions
Christian Urban <urbanc@in.tum.de>
parents: 882
diff changeset
   260
99e811fc1366 a few more lemmas...except supp of lambda-abstractions
Christian Urban <urbanc@in.tum.de>
parents: 882
diff changeset
   261
lemma var_supp1:
99e811fc1366 a few more lemmas...except supp of lambda-abstractions
Christian Urban <urbanc@in.tum.de>
parents: 882
diff changeset
   262
  shows "(supp (Var a)) = ((supp a)::name set)"
99e811fc1366 a few more lemmas...except supp of lambda-abstractions
Christian Urban <urbanc@in.tum.de>
parents: 882
diff changeset
   263
apply(simp add: supp_def pi_var1 var_inject)
99e811fc1366 a few more lemmas...except supp of lambda-abstractions
Christian Urban <urbanc@in.tum.de>
parents: 882
diff changeset
   264
done
99e811fc1366 a few more lemmas...except supp of lambda-abstractions
Christian Urban <urbanc@in.tum.de>
parents: 882
diff changeset
   265
99e811fc1366 a few more lemmas...except supp of lambda-abstractions
Christian Urban <urbanc@in.tum.de>
parents: 882
diff changeset
   266
lemma var_supp:
99e811fc1366 a few more lemmas...except supp of lambda-abstractions
Christian Urban <urbanc@in.tum.de>
parents: 882
diff changeset
   267
  shows "(supp (Var a)) = {a::name}"
99e811fc1366 a few more lemmas...except supp of lambda-abstractions
Christian Urban <urbanc@in.tum.de>
parents: 882
diff changeset
   268
using var_supp1
99e811fc1366 a few more lemmas...except supp of lambda-abstractions
Christian Urban <urbanc@in.tum.de>
parents: 882
diff changeset
   269
apply(simp add: supp_atm)
99e811fc1366 a few more lemmas...except supp of lambda-abstractions
Christian Urban <urbanc@in.tum.de>
parents: 882
diff changeset
   270
done
99e811fc1366 a few more lemmas...except supp of lambda-abstractions
Christian Urban <urbanc@in.tum.de>
parents: 882
diff changeset
   271
99e811fc1366 a few more lemmas...except supp of lambda-abstractions
Christian Urban <urbanc@in.tum.de>
parents: 882
diff changeset
   272
lemma app_supp:
99e811fc1366 a few more lemmas...except supp of lambda-abstractions
Christian Urban <urbanc@in.tum.de>
parents: 882
diff changeset
   273
  shows "supp (App t1 t2) = (supp t1) \<union> ((supp t2)::name set)"
99e811fc1366 a few more lemmas...except supp of lambda-abstractions
Christian Urban <urbanc@in.tum.de>
parents: 882
diff changeset
   274
apply(simp only: supp_def pi_app app_inject)
99e811fc1366 a few more lemmas...except supp of lambda-abstractions
Christian Urban <urbanc@in.tum.de>
parents: 882
diff changeset
   275
apply(simp add: Collect_imp_eq Collect_neg_eq)
99e811fc1366 a few more lemmas...except supp of lambda-abstractions
Christian Urban <urbanc@in.tum.de>
parents: 882
diff changeset
   276
done
99e811fc1366 a few more lemmas...except supp of lambda-abstractions
Christian Urban <urbanc@in.tum.de>
parents: 882
diff changeset
   277
99e811fc1366 a few more lemmas...except supp of lambda-abstractions
Christian Urban <urbanc@in.tum.de>
parents: 882
diff changeset
   278
lemma lam_supp:
99e811fc1366 a few more lemmas...except supp of lambda-abstractions
Christian Urban <urbanc@in.tum.de>
parents: 882
diff changeset
   279
  shows "supp (Lam x t) = ((supp ([x].t))::name set)"
99e811fc1366 a few more lemmas...except supp of lambda-abstractions
Christian Urban <urbanc@in.tum.de>
parents: 882
diff changeset
   280
apply(simp add: supp_def pi_lam)
99e811fc1366 a few more lemmas...except supp of lambda-abstractions
Christian Urban <urbanc@in.tum.de>
parents: 882
diff changeset
   281
sorry
99e811fc1366 a few more lemmas...except supp of lambda-abstractions
Christian Urban <urbanc@in.tum.de>
parents: 882
diff changeset
   282
804
Christian Urban <urbanc@in.tum.de>
parents: 767
diff changeset
   283
lemma lam_induct:
877
09a64cb04851 exported absrep_const for nitpick.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 876
diff changeset
   284
  "\<lbrakk>\<And>name. P (Var name);
804
Christian Urban <urbanc@in.tum.de>
parents: 767
diff changeset
   285
    \<And>lam1 lam2. \<lbrakk>P lam1; P lam2\<rbrakk> \<Longrightarrow> P (App lam1 lam2);
878
c3662f845129 setup for strong induction
Christian Urban <urbanc@in.tum.de>
parents: 877
diff changeset
   286
    \<And>name lam. P lam \<Longrightarrow> P (Lam name lam)\<rbrakk> 
c3662f845129 setup for strong induction
Christian Urban <urbanc@in.tum.de>
parents: 877
diff changeset
   287
    \<Longrightarrow> P lam"
804
Christian Urban <urbanc@in.tum.de>
parents: 767
diff changeset
   288
  by (lifting rlam.induct)
451
586e3dc4afdb Added 'TRY' to refl in clean_tac to get as far as possible. Removed unnecessary [quot_rsp] in FSet. Added necessary [quot_rsp] and one lifted thm in LamEx.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 445
diff changeset
   289
878
c3662f845129 setup for strong induction
Christian Urban <urbanc@in.tum.de>
parents: 877
diff changeset
   290
instance lam::pt_name
c3662f845129 setup for strong induction
Christian Urban <urbanc@in.tum.de>
parents: 877
diff changeset
   291
apply(default)
883
99e811fc1366 a few more lemmas...except supp of lambda-abstractions
Christian Urban <urbanc@in.tum.de>
parents: 882
diff changeset
   292
apply(induct_tac x rule: lam_induct)
99e811fc1366 a few more lemmas...except supp of lambda-abstractions
Christian Urban <urbanc@in.tum.de>
parents: 882
diff changeset
   293
apply(simp add: pi_var1)
99e811fc1366 a few more lemmas...except supp of lambda-abstractions
Christian Urban <urbanc@in.tum.de>
parents: 882
diff changeset
   294
apply(simp add: pi_app)
99e811fc1366 a few more lemmas...except supp of lambda-abstractions
Christian Urban <urbanc@in.tum.de>
parents: 882
diff changeset
   295
apply(simp add: pi_lam)
99e811fc1366 a few more lemmas...except supp of lambda-abstractions
Christian Urban <urbanc@in.tum.de>
parents: 882
diff changeset
   296
apply(induct_tac x rule: lam_induct)
99e811fc1366 a few more lemmas...except supp of lambda-abstractions
Christian Urban <urbanc@in.tum.de>
parents: 882
diff changeset
   297
apply(simp add: pi_var1 pt_name2)
99e811fc1366 a few more lemmas...except supp of lambda-abstractions
Christian Urban <urbanc@in.tum.de>
parents: 882
diff changeset
   298
apply(simp add: pi_app)
99e811fc1366 a few more lemmas...except supp of lambda-abstractions
Christian Urban <urbanc@in.tum.de>
parents: 882
diff changeset
   299
apply(simp add: pi_lam pt_name2)
99e811fc1366 a few more lemmas...except supp of lambda-abstractions
Christian Urban <urbanc@in.tum.de>
parents: 882
diff changeset
   300
apply(induct_tac x rule: lam_induct)
99e811fc1366 a few more lemmas...except supp of lambda-abstractions
Christian Urban <urbanc@in.tum.de>
parents: 882
diff changeset
   301
apply(simp add: pi_var1 pt_name3)
99e811fc1366 a few more lemmas...except supp of lambda-abstractions
Christian Urban <urbanc@in.tum.de>
parents: 882
diff changeset
   302
apply(simp add: pi_app)
99e811fc1366 a few more lemmas...except supp of lambda-abstractions
Christian Urban <urbanc@in.tum.de>
parents: 882
diff changeset
   303
apply(simp add: pi_lam pt_name3)
99e811fc1366 a few more lemmas...except supp of lambda-abstractions
Christian Urban <urbanc@in.tum.de>
parents: 882
diff changeset
   304
done
878
c3662f845129 setup for strong induction
Christian Urban <urbanc@in.tum.de>
parents: 877
diff changeset
   305
c3662f845129 setup for strong induction
Christian Urban <urbanc@in.tum.de>
parents: 877
diff changeset
   306
instance lam::fs_name
c3662f845129 setup for strong induction
Christian Urban <urbanc@in.tum.de>
parents: 877
diff changeset
   307
apply(default)
883
99e811fc1366 a few more lemmas...except supp of lambda-abstractions
Christian Urban <urbanc@in.tum.de>
parents: 882
diff changeset
   308
apply(induct_tac x rule: lam_induct)
99e811fc1366 a few more lemmas...except supp of lambda-abstractions
Christian Urban <urbanc@in.tum.de>
parents: 882
diff changeset
   309
apply(simp add: var_supp)
99e811fc1366 a few more lemmas...except supp of lambda-abstractions
Christian Urban <urbanc@in.tum.de>
parents: 882
diff changeset
   310
apply(simp add: app_supp)
877
09a64cb04851 exported absrep_const for nitpick.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 876
diff changeset
   311
sorry
09a64cb04851 exported absrep_const for nitpick.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 876
diff changeset
   312
881
2cc520457e37 nearly all of the proof
Christian Urban <urbanc@in.tum.de>
parents: 880
diff changeset
   313
lemma fresh_lam:
2cc520457e37 nearly all of the proof
Christian Urban <urbanc@in.tum.de>
parents: 880
diff changeset
   314
  "(a \<sharp> Lam b t) \<longleftrightarrow> (a = b) \<or> (a \<noteq> b \<and> a \<sharp> t)"
883
99e811fc1366 a few more lemmas...except supp of lambda-abstractions
Christian Urban <urbanc@in.tum.de>
parents: 882
diff changeset
   315
apply(simp add: fresh_def)
99e811fc1366 a few more lemmas...except supp of lambda-abstractions
Christian Urban <urbanc@in.tum.de>
parents: 882
diff changeset
   316
apply(simp add: lam_supp abs_supp)
99e811fc1366 a few more lemmas...except supp of lambda-abstractions
Christian Urban <urbanc@in.tum.de>
parents: 882
diff changeset
   317
apply(auto)
99e811fc1366 a few more lemmas...except supp of lambda-abstractions
Christian Urban <urbanc@in.tum.de>
parents: 882
diff changeset
   318
done
881
2cc520457e37 nearly all of the proof
Christian Urban <urbanc@in.tum.de>
parents: 880
diff changeset
   319
877
09a64cb04851 exported absrep_const for nitpick.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 876
diff changeset
   320
lemma lam_induct_strong:
878
c3662f845129 setup for strong induction
Christian Urban <urbanc@in.tum.de>
parents: 877
diff changeset
   321
  fixes a::"'a::fs_name"
879
f2a1ebba9bdc First subgoal.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 878
diff changeset
   322
  assumes a1: "\<And>name b. P b (Var name)"
f2a1ebba9bdc First subgoal.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 878
diff changeset
   323
  and     a2: "\<And>lam1 lam2 b. \<lbrakk>\<And>c. P c lam1; \<And>c. P c lam2\<rbrakk> \<Longrightarrow> P b (App lam1 lam2)"
881
2cc520457e37 nearly all of the proof
Christian Urban <urbanc@in.tum.de>
parents: 880
diff changeset
   324
  and     a3: "\<And>name lam b. \<lbrakk>\<And>c. P c lam; name \<sharp> b\<rbrakk> \<Longrightarrow> P b (Lam name lam)"
879
f2a1ebba9bdc First subgoal.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 878
diff changeset
   325
  shows "P a lam"
878
c3662f845129 setup for strong induction
Christian Urban <urbanc@in.tum.de>
parents: 877
diff changeset
   326
proof -
880
cd3f1409780a right generalisation
Christian Urban <urbanc@in.tum.de>
parents: 879
diff changeset
   327
  have "\<And>(pi::name prm) a. P a (pi \<bullet> lam)" 
878
c3662f845129 setup for strong induction
Christian Urban <urbanc@in.tum.de>
parents: 877
diff changeset
   328
  proof (induct lam rule: lam_induct)
c3662f845129 setup for strong induction
Christian Urban <urbanc@in.tum.de>
parents: 877
diff changeset
   329
    case (1 name pi)
879
f2a1ebba9bdc First subgoal.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 878
diff changeset
   330
    show "P a (pi \<bullet> Var name)"
f2a1ebba9bdc First subgoal.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 878
diff changeset
   331
      apply (simp only: pi_var1)
f2a1ebba9bdc First subgoal.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 878
diff changeset
   332
      apply (rule a1)
f2a1ebba9bdc First subgoal.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 878
diff changeset
   333
      done
878
c3662f845129 setup for strong induction
Christian Urban <urbanc@in.tum.de>
parents: 877
diff changeset
   334
  next
c3662f845129 setup for strong induction
Christian Urban <urbanc@in.tum.de>
parents: 877
diff changeset
   335
    case (2 lam1 lam2 pi)
880
cd3f1409780a right generalisation
Christian Urban <urbanc@in.tum.de>
parents: 879
diff changeset
   336
    have b1: "\<And>(pi::name prm) a. P a (pi \<bullet> lam1)" by fact
cd3f1409780a right generalisation
Christian Urban <urbanc@in.tum.de>
parents: 879
diff changeset
   337
    have b2: "\<And>(pi::name prm) a. P a (pi \<bullet> lam2)" by fact
879
f2a1ebba9bdc First subgoal.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 878
diff changeset
   338
    show "P a (pi \<bullet> App lam1 lam2)"
f2a1ebba9bdc First subgoal.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 878
diff changeset
   339
      apply (simp only: pi_app)
f2a1ebba9bdc First subgoal.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 878
diff changeset
   340
      apply (rule a2)
880
cd3f1409780a right generalisation
Christian Urban <urbanc@in.tum.de>
parents: 879
diff changeset
   341
      apply (rule b1)
cd3f1409780a right generalisation
Christian Urban <urbanc@in.tum.de>
parents: 879
diff changeset
   342
      apply (rule b2)
cd3f1409780a right generalisation
Christian Urban <urbanc@in.tum.de>
parents: 879
diff changeset
   343
      done
878
c3662f845129 setup for strong induction
Christian Urban <urbanc@in.tum.de>
parents: 877
diff changeset
   344
  next
881
2cc520457e37 nearly all of the proof
Christian Urban <urbanc@in.tum.de>
parents: 880
diff changeset
   345
    case (3 name lam pi a)
2cc520457e37 nearly all of the proof
Christian Urban <urbanc@in.tum.de>
parents: 880
diff changeset
   346
    have b: "\<And>(pi::name prm) a. P a (pi \<bullet> lam)" by fact
882
6a8858ba01f6 removed one sorry
Christian Urban <urbanc@in.tum.de>
parents: 881
diff changeset
   347
    obtain c::name where fr: "c\<sharp>(a, pi\<bullet>name, pi\<bullet>lam)"
6a8858ba01f6 removed one sorry
Christian Urban <urbanc@in.tum.de>
parents: 881
diff changeset
   348
      apply(rule exists_fresh[of "(a, pi\<bullet>name, pi\<bullet>lam)"])
6a8858ba01f6 removed one sorry
Christian Urban <urbanc@in.tum.de>
parents: 881
diff changeset
   349
      apply(simp_all add: fs_name1)
6a8858ba01f6 removed one sorry
Christian Urban <urbanc@in.tum.de>
parents: 881
diff changeset
   350
      done
881
2cc520457e37 nearly all of the proof
Christian Urban <urbanc@in.tum.de>
parents: 880
diff changeset
   351
    from b fr have p: "P a (Lam c (([(c, pi\<bullet>name)]@pi)\<bullet>lam))" 
2cc520457e37 nearly all of the proof
Christian Urban <urbanc@in.tum.de>
parents: 880
diff changeset
   352
      apply -
2cc520457e37 nearly all of the proof
Christian Urban <urbanc@in.tum.de>
parents: 880
diff changeset
   353
      apply(rule a3)
2cc520457e37 nearly all of the proof
Christian Urban <urbanc@in.tum.de>
parents: 880
diff changeset
   354
      apply(blast)
2cc520457e37 nearly all of the proof
Christian Urban <urbanc@in.tum.de>
parents: 880
diff changeset
   355
      apply(simp)
2cc520457e37 nearly all of the proof
Christian Urban <urbanc@in.tum.de>
parents: 880
diff changeset
   356
      done
2cc520457e37 nearly all of the proof
Christian Urban <urbanc@in.tum.de>
parents: 880
diff changeset
   357
    have eq: "[(c, pi\<bullet>name)] \<bullet> Lam (pi \<bullet> name) (pi \<bullet> lam) = Lam (pi \<bullet> name) (pi \<bullet> lam)"
2cc520457e37 nearly all of the proof
Christian Urban <urbanc@in.tum.de>
parents: 880
diff changeset
   358
      apply(rule perm_fresh_fresh)
2cc520457e37 nearly all of the proof
Christian Urban <urbanc@in.tum.de>
parents: 880
diff changeset
   359
      using fr
2cc520457e37 nearly all of the proof
Christian Urban <urbanc@in.tum.de>
parents: 880
diff changeset
   360
      apply(simp add: fresh_lam)
2cc520457e37 nearly all of the proof
Christian Urban <urbanc@in.tum.de>
parents: 880
diff changeset
   361
      apply(simp add: fresh_lam)
2cc520457e37 nearly all of the proof
Christian Urban <urbanc@in.tum.de>
parents: 880
diff changeset
   362
      done
879
f2a1ebba9bdc First subgoal.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 878
diff changeset
   363
    show "P a (pi \<bullet> Lam name lam)" 
f2a1ebba9bdc First subgoal.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 878
diff changeset
   364
      apply (simp add: pi_lam)
881
2cc520457e37 nearly all of the proof
Christian Urban <urbanc@in.tum.de>
parents: 880
diff changeset
   365
      apply(subst eq[symmetric])
2cc520457e37 nearly all of the proof
Christian Urban <urbanc@in.tum.de>
parents: 880
diff changeset
   366
      using p
2cc520457e37 nearly all of the proof
Christian Urban <urbanc@in.tum.de>
parents: 880
diff changeset
   367
      apply(simp only: pi_lam pt_name2 swap_simps)
2cc520457e37 nearly all of the proof
Christian Urban <urbanc@in.tum.de>
parents: 880
diff changeset
   368
      done
878
c3662f845129 setup for strong induction
Christian Urban <urbanc@in.tum.de>
parents: 877
diff changeset
   369
  qed
c3662f845129 setup for strong induction
Christian Urban <urbanc@in.tum.de>
parents: 877
diff changeset
   370
  then have "P a (([]::name prm) \<bullet> lam)" by blast
c3662f845129 setup for strong induction
Christian Urban <urbanc@in.tum.de>
parents: 877
diff changeset
   371
  then show "P a lam" by simp 
c3662f845129 setup for strong induction
Christian Urban <urbanc@in.tum.de>
parents: 877
diff changeset
   372
qed
877
09a64cb04851 exported absrep_const for nitpick.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 876
diff changeset
   373
09a64cb04851 exported absrep_const for nitpick.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 876
diff changeset
   374
249
7dec34d12328 added some facts about fresh and support of lam
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
   375
lemma var_fresh:
7dec34d12328 added some facts about fresh and support of lam
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
   376
  fixes a::"name"
804
Christian Urban <urbanc@in.tum.de>
parents: 767
diff changeset
   377
  shows "(a \<sharp> (Var b)) = (a \<sharp> b)"
249
7dec34d12328 added some facts about fresh and support of lam
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
   378
  apply(simp add: fresh_def)
884
e49c6b6f37f4 tuned quotient_def.ML and cleaned somewhat LamEx.thy
Christian Urban <urbanc@in.tum.de>
parents: 883
diff changeset
   379
  apply(simp add: var_supp1)
249
7dec34d12328 added some facts about fresh and support of lam
Christian Urban <urbanc@in.tum.de>
parents: 247
diff changeset
   380
  done
247
e83a6e452843 Lemmas about fv.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 246
diff changeset
   381
891
7bac7dffadeb hom lifted to hom', so it is true. Infrastructure for partially regularized quantifiers. Nicer errors for regularize.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 890
diff changeset
   382
(* lemma hom_reg: *)
887
d2660637e764 Incorrect version of the homomorphism lemma
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 884
diff changeset
   383
895
92c43c96027e added a partial proof under which conditions rlam_rec Respects alpha...I guess something like this is true; this means the Hom lemmas need to have preconditions
Christian Urban <urbanc@in.tum.de>
parents: 894
diff changeset
   384
lemma rlam_rec_eqvt:
92c43c96027e added a partial proof under which conditions rlam_rec Respects alpha...I guess something like this is true; this means the Hom lemmas need to have preconditions
Christian Urban <urbanc@in.tum.de>
parents: 894
diff changeset
   385
  fixes pi::"name prm"
92c43c96027e added a partial proof under which conditions rlam_rec Respects alpha...I guess something like this is true; this means the Hom lemmas need to have preconditions
Christian Urban <urbanc@in.tum.de>
parents: 894
diff changeset
   386
  and   f1::"name \<Rightarrow> ('a::pt_name)"
92c43c96027e added a partial proof under which conditions rlam_rec Respects alpha...I guess something like this is true; this means the Hom lemmas need to have preconditions
Christian Urban <urbanc@in.tum.de>
parents: 894
diff changeset
   387
  shows "(pi\<bullet>rlam_rec f1 f2 f3 t) = rlam_rec (pi\<bullet>f1) (pi\<bullet>f2) (pi\<bullet>f3) (pi\<bullet>t)"
92c43c96027e added a partial proof under which conditions rlam_rec Respects alpha...I guess something like this is true; this means the Hom lemmas need to have preconditions
Christian Urban <urbanc@in.tum.de>
parents: 894
diff changeset
   388
apply(induct t)
92c43c96027e added a partial proof under which conditions rlam_rec Respects alpha...I guess something like this is true; this means the Hom lemmas need to have preconditions
Christian Urban <urbanc@in.tum.de>
parents: 894
diff changeset
   389
apply(simp_all)
92c43c96027e added a partial proof under which conditions rlam_rec Respects alpha...I guess something like this is true; this means the Hom lemmas need to have preconditions
Christian Urban <urbanc@in.tum.de>
parents: 894
diff changeset
   390
apply(simp add: perm_fun_def)
92c43c96027e added a partial proof under which conditions rlam_rec Respects alpha...I guess something like this is true; this means the Hom lemmas need to have preconditions
Christian Urban <urbanc@in.tum.de>
parents: 894
diff changeset
   391
apply(perm_simp)
92c43c96027e added a partial proof under which conditions rlam_rec Respects alpha...I guess something like this is true; this means the Hom lemmas need to have preconditions
Christian Urban <urbanc@in.tum.de>
parents: 894
diff changeset
   392
apply(subst pt_fun_app_eq[OF pt_name_inst at_name_inst])
92c43c96027e added a partial proof under which conditions rlam_rec Respects alpha...I guess something like this is true; this means the Hom lemmas need to have preconditions
Christian Urban <urbanc@in.tum.de>
parents: 894
diff changeset
   393
back
92c43c96027e added a partial proof under which conditions rlam_rec Respects alpha...I guess something like this is true; this means the Hom lemmas need to have preconditions
Christian Urban <urbanc@in.tum.de>
parents: 894
diff changeset
   394
apply(subst pt_fun_app_eq[OF pt_name_inst at_name_inst])
92c43c96027e added a partial proof under which conditions rlam_rec Respects alpha...I guess something like this is true; this means the Hom lemmas need to have preconditions
Christian Urban <urbanc@in.tum.de>
parents: 894
diff changeset
   395
apply(subst pt_fun_app_eq[OF pt_name_inst at_name_inst])
92c43c96027e added a partial proof under which conditions rlam_rec Respects alpha...I guess something like this is true; this means the Hom lemmas need to have preconditions
Christian Urban <urbanc@in.tum.de>
parents: 894
diff changeset
   396
apply(subst pt_fun_app_eq[OF pt_name_inst at_name_inst])
92c43c96027e added a partial proof under which conditions rlam_rec Respects alpha...I guess something like this is true; this means the Hom lemmas need to have preconditions
Christian Urban <urbanc@in.tum.de>
parents: 894
diff changeset
   397
apply(simp)
92c43c96027e added a partial proof under which conditions rlam_rec Respects alpha...I guess something like this is true; this means the Hom lemmas need to have preconditions
Christian Urban <urbanc@in.tum.de>
parents: 894
diff changeset
   398
apply(subst pt_fun_app_eq[OF pt_name_inst at_name_inst])
92c43c96027e added a partial proof under which conditions rlam_rec Respects alpha...I guess something like this is true; this means the Hom lemmas need to have preconditions
Christian Urban <urbanc@in.tum.de>
parents: 894
diff changeset
   399
back
92c43c96027e added a partial proof under which conditions rlam_rec Respects alpha...I guess something like this is true; this means the Hom lemmas need to have preconditions
Christian Urban <urbanc@in.tum.de>
parents: 894
diff changeset
   400
apply(subst pt_fun_app_eq[OF pt_name_inst at_name_inst])
92c43c96027e added a partial proof under which conditions rlam_rec Respects alpha...I guess something like this is true; this means the Hom lemmas need to have preconditions
Christian Urban <urbanc@in.tum.de>
parents: 894
diff changeset
   401
apply(subst pt_fun_app_eq[OF pt_name_inst at_name_inst])
92c43c96027e added a partial proof under which conditions rlam_rec Respects alpha...I guess something like this is true; this means the Hom lemmas need to have preconditions
Christian Urban <urbanc@in.tum.de>
parents: 894
diff changeset
   402
apply(simp)
92c43c96027e added a partial proof under which conditions rlam_rec Respects alpha...I guess something like this is true; this means the Hom lemmas need to have preconditions
Christian Urban <urbanc@in.tum.de>
parents: 894
diff changeset
   403
done
92c43c96027e added a partial proof under which conditions rlam_rec Respects alpha...I guess something like this is true; this means the Hom lemmas need to have preconditions
Christian Urban <urbanc@in.tum.de>
parents: 894
diff changeset
   404
 
92c43c96027e added a partial proof under which conditions rlam_rec Respects alpha...I guess something like this is true; this means the Hom lemmas need to have preconditions
Christian Urban <urbanc@in.tum.de>
parents: 894
diff changeset
   405
92c43c96027e added a partial proof under which conditions rlam_rec Respects alpha...I guess something like this is true; this means the Hom lemmas need to have preconditions
Christian Urban <urbanc@in.tum.de>
parents: 894
diff changeset
   406
lemma rlam_rec_respects:
92c43c96027e added a partial proof under which conditions rlam_rec Respects alpha...I guess something like this is true; this means the Hom lemmas need to have preconditions
Christian Urban <urbanc@in.tum.de>
parents: 894
diff changeset
   407
  assumes f1: "f_var \<in> Respects (op= ===> op=)"
92c43c96027e added a partial proof under which conditions rlam_rec Respects alpha...I guess something like this is true; this means the Hom lemmas need to have preconditions
Christian Urban <urbanc@in.tum.de>
parents: 894
diff changeset
   408
  and     f2: "f_app \<in> Respects (alpha ===> alpha ===> op= ===> op= ===> op=)"
92c43c96027e added a partial proof under which conditions rlam_rec Respects alpha...I guess something like this is true; this means the Hom lemmas need to have preconditions
Christian Urban <urbanc@in.tum.de>
parents: 894
diff changeset
   409
  and     f3: "f_lam \<in> Respects (op= ===> alpha ===> op= ===> op=)"
92c43c96027e added a partial proof under which conditions rlam_rec Respects alpha...I guess something like this is true; this means the Hom lemmas need to have preconditions
Christian Urban <urbanc@in.tum.de>
parents: 894
diff changeset
   410
  shows "rlam_rec f_var f_app f_lam \<in> Respects (alpha ===> op =)"
92c43c96027e added a partial proof under which conditions rlam_rec Respects alpha...I guess something like this is true; this means the Hom lemmas need to have preconditions
Christian Urban <urbanc@in.tum.de>
parents: 894
diff changeset
   411
apply(simp add: mem_def)
92c43c96027e added a partial proof under which conditions rlam_rec Respects alpha...I guess something like this is true; this means the Hom lemmas need to have preconditions
Christian Urban <urbanc@in.tum.de>
parents: 894
diff changeset
   412
apply(simp add: Respects_def)
92c43c96027e added a partial proof under which conditions rlam_rec Respects alpha...I guess something like this is true; this means the Hom lemmas need to have preconditions
Christian Urban <urbanc@in.tum.de>
parents: 894
diff changeset
   413
apply(rule allI)
92c43c96027e added a partial proof under which conditions rlam_rec Respects alpha...I guess something like this is true; this means the Hom lemmas need to have preconditions
Christian Urban <urbanc@in.tum.de>
parents: 894
diff changeset
   414
apply(rule allI)
92c43c96027e added a partial proof under which conditions rlam_rec Respects alpha...I guess something like this is true; this means the Hom lemmas need to have preconditions
Christian Urban <urbanc@in.tum.de>
parents: 894
diff changeset
   415
apply(rule impI)
92c43c96027e added a partial proof under which conditions rlam_rec Respects alpha...I guess something like this is true; this means the Hom lemmas need to have preconditions
Christian Urban <urbanc@in.tum.de>
parents: 894
diff changeset
   416
apply(erule alpha.induct)
92c43c96027e added a partial proof under which conditions rlam_rec Respects alpha...I guess something like this is true; this means the Hom lemmas need to have preconditions
Christian Urban <urbanc@in.tum.de>
parents: 894
diff changeset
   417
apply(simp)
92c43c96027e added a partial proof under which conditions rlam_rec Respects alpha...I guess something like this is true; this means the Hom lemmas need to have preconditions
Christian Urban <urbanc@in.tum.de>
parents: 894
diff changeset
   418
apply(simp)
92c43c96027e added a partial proof under which conditions rlam_rec Respects alpha...I guess something like this is true; this means the Hom lemmas need to have preconditions
Christian Urban <urbanc@in.tum.de>
parents: 894
diff changeset
   419
using f2
92c43c96027e added a partial proof under which conditions rlam_rec Respects alpha...I guess something like this is true; this means the Hom lemmas need to have preconditions
Christian Urban <urbanc@in.tum.de>
parents: 894
diff changeset
   420
apply(simp add: mem_def)
92c43c96027e added a partial proof under which conditions rlam_rec Respects alpha...I guess something like this is true; this means the Hom lemmas need to have preconditions
Christian Urban <urbanc@in.tum.de>
parents: 894
diff changeset
   421
apply(simp add: Respects_def)
92c43c96027e added a partial proof under which conditions rlam_rec Respects alpha...I guess something like this is true; this means the Hom lemmas need to have preconditions
Christian Urban <urbanc@in.tum.de>
parents: 894
diff changeset
   422
using f3[simplified mem_def Respects_def]
92c43c96027e added a partial proof under which conditions rlam_rec Respects alpha...I guess something like this is true; this means the Hom lemmas need to have preconditions
Christian Urban <urbanc@in.tum.de>
parents: 894
diff changeset
   423
apply(simp)
92c43c96027e added a partial proof under which conditions rlam_rec Respects alpha...I guess something like this is true; this means the Hom lemmas need to have preconditions
Christian Urban <urbanc@in.tum.de>
parents: 894
diff changeset
   424
apply(case_tac "a=b")
92c43c96027e added a partial proof under which conditions rlam_rec Respects alpha...I guess something like this is true; this means the Hom lemmas need to have preconditions
Christian Urban <urbanc@in.tum.de>
parents: 894
diff changeset
   425
apply(clarify)
92c43c96027e added a partial proof under which conditions rlam_rec Respects alpha...I guess something like this is true; this means the Hom lemmas need to have preconditions
Christian Urban <urbanc@in.tum.de>
parents: 894
diff changeset
   426
apply(simp)
92c43c96027e added a partial proof under which conditions rlam_rec Respects alpha...I guess something like this is true; this means the Hom lemmas need to have preconditions
Christian Urban <urbanc@in.tum.de>
parents: 894
diff changeset
   427
apply(subst pt_swap_bij'')
92c43c96027e added a partial proof under which conditions rlam_rec Respects alpha...I guess something like this is true; this means the Hom lemmas need to have preconditions
Christian Urban <urbanc@in.tum.de>
parents: 894
diff changeset
   428
apply(rule pt_name_inst)
92c43c96027e added a partial proof under which conditions rlam_rec Respects alpha...I guess something like this is true; this means the Hom lemmas need to have preconditions
Christian Urban <urbanc@in.tum.de>
parents: 894
diff changeset
   429
apply(rule at_name_inst)
92c43c96027e added a partial proof under which conditions rlam_rec Respects alpha...I guess something like this is true; this means the Hom lemmas need to have preconditions
Christian Urban <urbanc@in.tum.de>
parents: 894
diff changeset
   430
apply(subst pt_swap_bij'')
92c43c96027e added a partial proof under which conditions rlam_rec Respects alpha...I guess something like this is true; this means the Hom lemmas need to have preconditions
Christian Urban <urbanc@in.tum.de>
parents: 894
diff changeset
   431
apply(rule pt_name_inst)
92c43c96027e added a partial proof under which conditions rlam_rec Respects alpha...I guess something like this is true; this means the Hom lemmas need to have preconditions
Christian Urban <urbanc@in.tum.de>
parents: 894
diff changeset
   432
apply(rule at_name_inst)
92c43c96027e added a partial proof under which conditions rlam_rec Respects alpha...I guess something like this is true; this means the Hom lemmas need to have preconditions
Christian Urban <urbanc@in.tum.de>
parents: 894
diff changeset
   433
apply(simp)
92c43c96027e added a partial proof under which conditions rlam_rec Respects alpha...I guess something like this is true; this means the Hom lemmas need to have preconditions
Christian Urban <urbanc@in.tum.de>
parents: 894
diff changeset
   434
apply(generate_fresh "name")
92c43c96027e added a partial proof under which conditions rlam_rec Respects alpha...I guess something like this is true; this means the Hom lemmas need to have preconditions
Christian Urban <urbanc@in.tum.de>
parents: 894
diff changeset
   435
(* probably true *)
92c43c96027e added a partial proof under which conditions rlam_rec Respects alpha...I guess something like this is true; this means the Hom lemmas need to have preconditions
Christian Urban <urbanc@in.tum.de>
parents: 894
diff changeset
   436
sorry
92c43c96027e added a partial proof under which conditions rlam_rec Respects alpha...I guess something like this is true; this means the Hom lemmas need to have preconditions
Christian Urban <urbanc@in.tum.de>
parents: 894
diff changeset
   437
894
1d80641a4302 tried to witness the hom-lemma with the recursion combinator from rlam....does not work yet completely
Christian Urban <urbanc@in.tum.de>
parents: 891
diff changeset
   438
lemma hom: 
1d80641a4302 tried to witness the hom-lemma with the recursion combinator from rlam....does not work yet completely
Christian Urban <urbanc@in.tum.de>
parents: 891
diff changeset
   439
  "\<exists>hom\<in>Respects (alpha ===> op =). 
1d80641a4302 tried to witness the hom-lemma with the recursion combinator from rlam....does not work yet completely
Christian Urban <urbanc@in.tum.de>
parents: 891
diff changeset
   440
    ((\<forall>x. hom (rVar x) = f_var x) \<and>
1d80641a4302 tried to witness the hom-lemma with the recursion combinator from rlam....does not work yet completely
Christian Urban <urbanc@in.tum.de>
parents: 891
diff changeset
   441
     (\<forall>l r. hom (rApp l r) = f_app l r (hom l) (hom r)) \<and>
1d80641a4302 tried to witness the hom-lemma with the recursion combinator from rlam....does not work yet completely
Christian Urban <urbanc@in.tum.de>
parents: 891
diff changeset
   442
     (\<forall>x a. hom (rLam a x) = f_lam (\<lambda>b. ([(a,b)]\<bullet> x)) (\<lambda>b. hom ([(a,b)] \<bullet> x))))"
1d80641a4302 tried to witness the hom-lemma with the recursion combinator from rlam....does not work yet completely
Christian Urban <urbanc@in.tum.de>
parents: 891
diff changeset
   443
apply(rule_tac x="rlam_rec f_var f_app XX" in bexI)
1d80641a4302 tried to witness the hom-lemma with the recursion combinator from rlam....does not work yet completely
Christian Urban <urbanc@in.tum.de>
parents: 891
diff changeset
   444
apply(rule conjI)
1d80641a4302 tried to witness the hom-lemma with the recursion combinator from rlam....does not work yet completely
Christian Urban <urbanc@in.tum.de>
parents: 891
diff changeset
   445
apply(simp)
1d80641a4302 tried to witness the hom-lemma with the recursion combinator from rlam....does not work yet completely
Christian Urban <urbanc@in.tum.de>
parents: 891
diff changeset
   446
apply(rule conjI)
1d80641a4302 tried to witness the hom-lemma with the recursion combinator from rlam....does not work yet completely
Christian Urban <urbanc@in.tum.de>
parents: 891
diff changeset
   447
apply(simp)
1d80641a4302 tried to witness the hom-lemma with the recursion combinator from rlam....does not work yet completely
Christian Urban <urbanc@in.tum.de>
parents: 891
diff changeset
   448
apply(simp)
889
cff21786d952 Appropriate respects and a statement of the lifted hom lemma
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 888
diff changeset
   449
sorry
cff21786d952 Appropriate respects and a statement of the lifted hom lemma
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 888
diff changeset
   450
891
7bac7dffadeb hom lifted to hom', so it is true. Infrastructure for partially regularized quantifiers. Nicer errors for regularize.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 890
diff changeset
   451
lemma hom_reg:"
7bac7dffadeb hom lifted to hom', so it is true. Infrastructure for partially regularized quantifiers. Nicer errors for regularize.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 890
diff changeset
   452
(\<exists>hom\<in>Respects (alpha ===> op =).
7bac7dffadeb hom lifted to hom', so it is true. Infrastructure for partially regularized quantifiers. Nicer errors for regularize.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 890
diff changeset
   453
 (\<forall>x. hom (rVar x) = f_var x) \<and>
7bac7dffadeb hom lifted to hom', so it is true. Infrastructure for partially regularized quantifiers. Nicer errors for regularize.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 890
diff changeset
   454
 (\<forall>l r. hom (rApp l r) = f_app l r (hom l) (hom r)) \<and>
7bac7dffadeb hom lifted to hom', so it is true. Infrastructure for partially regularized quantifiers. Nicer errors for regularize.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 890
diff changeset
   455
 (\<forall>xa a. hom (rLam a xa) = f_lam (\<lambda>b. [(a, b)] \<bullet> xa) (\<lambda>b. hom ([(a, b)] \<bullet> xa)))) \<longrightarrow>
7bac7dffadeb hom lifted to hom', so it is true. Infrastructure for partially regularized quantifiers. Nicer errors for regularize.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 890
diff changeset
   456
(\<exists>hom.
7bac7dffadeb hom lifted to hom', so it is true. Infrastructure for partially regularized quantifiers. Nicer errors for regularize.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 890
diff changeset
   457
 (\<forall>x. hom (rVar x) = f_var x) \<and>
7bac7dffadeb hom lifted to hom', so it is true. Infrastructure for partially regularized quantifiers. Nicer errors for regularize.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 890
diff changeset
   458
 (\<forall>l r. hom (rApp l r) = f_app l r (hom l) (hom r)) \<and>
7bac7dffadeb hom lifted to hom', so it is true. Infrastructure for partially regularized quantifiers. Nicer errors for regularize.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 890
diff changeset
   459
 (\<forall>xa a. hom (rLam a xa) = f_lam (\<lambda>b. [(a, b)] \<bullet> xa) (\<lambda>b. hom ([(a, b)] \<bullet> xa))))"
7bac7dffadeb hom lifted to hom', so it is true. Infrastructure for partially regularized quantifiers. Nicer errors for regularize.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 890
diff changeset
   460
by(regularize)
7bac7dffadeb hom lifted to hom', so it is true. Infrastructure for partially regularized quantifiers. Nicer errors for regularize.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 890
diff changeset
   461
7bac7dffadeb hom lifted to hom', so it is true. Infrastructure for partially regularized quantifiers. Nicer errors for regularize.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 890
diff changeset
   462
lemma hom_pre:"
7bac7dffadeb hom lifted to hom', so it is true. Infrastructure for partially regularized quantifiers. Nicer errors for regularize.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 890
diff changeset
   463
(\<exists>hom.
7bac7dffadeb hom lifted to hom', so it is true. Infrastructure for partially regularized quantifiers. Nicer errors for regularize.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 890
diff changeset
   464
 (\<forall>x. hom (rVar x) = f_var x) \<and>
7bac7dffadeb hom lifted to hom', so it is true. Infrastructure for partially regularized quantifiers. Nicer errors for regularize.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 890
diff changeset
   465
 (\<forall>l r. hom (rApp l r) = f_app l r (hom l) (hom r)) \<and>
7bac7dffadeb hom lifted to hom', so it is true. Infrastructure for partially regularized quantifiers. Nicer errors for regularize.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 890
diff changeset
   466
 (\<forall>xa a. hom (rLam a xa) = f_lam (\<lambda>b. [(a, b)] \<bullet> xa) (\<lambda>b. hom ([(a, b)] \<bullet> xa))))"
7bac7dffadeb hom lifted to hom', so it is true. Infrastructure for partially regularized quantifiers. Nicer errors for regularize.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 890
diff changeset
   467
apply (rule impE[OF hom_reg])
7bac7dffadeb hom lifted to hom', so it is true. Infrastructure for partially regularized quantifiers. Nicer errors for regularize.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 890
diff changeset
   468
apply (rule hom)
7bac7dffadeb hom lifted to hom', so it is true. Infrastructure for partially regularized quantifiers. Nicer errors for regularize.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 890
diff changeset
   469
apply (assumption)
7bac7dffadeb hom lifted to hom', so it is true. Infrastructure for partially regularized quantifiers. Nicer errors for regularize.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 890
diff changeset
   470
done
7bac7dffadeb hom lifted to hom', so it is true. Infrastructure for partially regularized quantifiers. Nicer errors for regularize.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 890
diff changeset
   471
889
cff21786d952 Appropriate respects and a statement of the lifted hom lemma
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 888
diff changeset
   472
lemma hom':
894
1d80641a4302 tried to witness the hom-lemma with the recursion combinator from rlam....does not work yet completely
Christian Urban <urbanc@in.tum.de>
parents: 891
diff changeset
   473
"\<exists>hom. 
1d80641a4302 tried to witness the hom-lemma with the recursion combinator from rlam....does not work yet completely
Christian Urban <urbanc@in.tum.de>
parents: 891
diff changeset
   474
  ((\<forall>x. hom (Var x) = f_var x) \<and>
1d80641a4302 tried to witness the hom-lemma with the recursion combinator from rlam....does not work yet completely
Christian Urban <urbanc@in.tum.de>
parents: 891
diff changeset
   475
   (\<forall>l r. hom (App l r) = f_app l r (hom l) (hom r)) \<and>
1d80641a4302 tried to witness the hom-lemma with the recursion combinator from rlam....does not work yet completely
Christian Urban <urbanc@in.tum.de>
parents: 891
diff changeset
   476
   (\<forall>x a. hom (Lam a x) = f_lam (\<lambda>b. ([(a,b)] \<bullet> x)) (\<lambda>b. hom ([(a,b)] \<bullet> x))))"
891
7bac7dffadeb hom lifted to hom', so it is true. Infrastructure for partially regularized quantifiers. Nicer errors for regularize.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 890
diff changeset
   477
apply (lifting hom)
7bac7dffadeb hom lifted to hom', so it is true. Infrastructure for partially regularized quantifiers. Nicer errors for regularize.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 890
diff changeset
   478
done
890
0f920b62fb7b slight tuning of relation_error
Christian Urban <urbanc@in.tum.de>
parents: 889
diff changeset
   479
891
7bac7dffadeb hom lifted to hom', so it is true. Infrastructure for partially regularized quantifiers. Nicer errors for regularize.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 890
diff changeset
   480
thm bex_reg_eqv_range[OF identity_equivp, of "alpha"]
7bac7dffadeb hom lifted to hom', so it is true. Infrastructure for partially regularized quantifiers. Nicer errors for regularize.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 890
diff changeset
   481
887
d2660637e764 Incorrect version of the homomorphism lemma
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 884
diff changeset
   482
663
0dd10a900cae Different syntax for definitions that allows overloading and retrieving of definitions by matching whole constants.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 636
diff changeset
   483
end
891
7bac7dffadeb hom lifted to hom', so it is true. Infrastructure for partially regularized quantifiers. Nicer errors for regularize.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 890
diff changeset
   484