Nominal/Term1.thy
author Cezary Kaliszyk <kaliszyk@in.tum.de>
Thu, 11 Mar 2010 16:15:29 +0100
changeset 1425 112f998d2953
parent 1400 10eca65a8d03
child 1428 4029105011ca
permissions -rw-r--r--
Fixes for term1 for new alpha. Still not able to show support equations.
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
1270
8c3cf9f4f5f2 Split Terms into separate files and add them to tests.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
     1
theory Term1
8c3cf9f4f5f2 Split Terms into separate files and add them to tests.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
     2
imports "Nominal2_Atoms" "Nominal2_Eqvt" "Nominal2_Supp" "Abs" "Perm" "Fv" "Rsp" "../Attic/Prove"
8c3cf9f4f5f2 Split Terms into separate files and add them to tests.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
     3
begin
8c3cf9f4f5f2 Split Terms into separate files and add them to tests.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
     4
8c3cf9f4f5f2 Split Terms into separate files and add them to tests.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
     5
atom_decl name
8c3cf9f4f5f2 Split Terms into separate files and add them to tests.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
     6
8c3cf9f4f5f2 Split Terms into separate files and add them to tests.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
     7
section {*** lets with binding patterns ***}
8c3cf9f4f5f2 Split Terms into separate files and add them to tests.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
     8
8c3cf9f4f5f2 Split Terms into separate files and add them to tests.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
     9
datatype rtrm1 =
8c3cf9f4f5f2 Split Terms into separate files and add them to tests.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    10
  rVr1 "name"
8c3cf9f4f5f2 Split Terms into separate files and add them to tests.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    11
| rAp1 "rtrm1" "rtrm1"
8c3cf9f4f5f2 Split Terms into separate files and add them to tests.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    12
| rLm1 "name" "rtrm1"        --"name is bound in trm1"
8c3cf9f4f5f2 Split Terms into separate files and add them to tests.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    13
| rLt1 "bp" "rtrm1" "rtrm1"   --"all variables in bp are bound in the 2nd trm1"
8c3cf9f4f5f2 Split Terms into separate files and add them to tests.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    14
and bp =
8c3cf9f4f5f2 Split Terms into separate files and add them to tests.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    15
  BUnit
8c3cf9f4f5f2 Split Terms into separate files and add them to tests.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    16
| BVr "name"
8c3cf9f4f5f2 Split Terms into separate files and add them to tests.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    17
| BPr "bp" "bp"
8c3cf9f4f5f2 Split Terms into separate files and add them to tests.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    18
8c3cf9f4f5f2 Split Terms into separate files and add them to tests.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    19
print_theorems
8c3cf9f4f5f2 Split Terms into separate files and add them to tests.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    20
8c3cf9f4f5f2 Split Terms into separate files and add them to tests.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    21
(* to be given by the user *)
8c3cf9f4f5f2 Split Terms into separate files and add them to tests.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    22
8c3cf9f4f5f2 Split Terms into separate files and add them to tests.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    23
primrec 
8c3cf9f4f5f2 Split Terms into separate files and add them to tests.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    24
  bv1
8c3cf9f4f5f2 Split Terms into separate files and add them to tests.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    25
where
8c3cf9f4f5f2 Split Terms into separate files and add them to tests.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    26
  "bv1 (BUnit) = {}"
8c3cf9f4f5f2 Split Terms into separate files and add them to tests.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    27
| "bv1 (BVr x) = {atom x}"
8c3cf9f4f5f2 Split Terms into separate files and add them to tests.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    28
| "bv1 (BPr bp1 bp2) = (bv1 bp1) \<union> (bv1 bp2)"
8c3cf9f4f5f2 Split Terms into separate files and add them to tests.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    29
1277
6eacf60ce41d Permutation and FV_Alpha interface change.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1274
diff changeset
    30
setup {* snd o define_raw_perms (Datatype.the_info @{theory} "Term1.rtrm1") 2 *}
1270
8c3cf9f4f5f2 Split Terms into separate files and add them to tests.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    31
thm permute_rtrm1_permute_bp.simps
8c3cf9f4f5f2 Split Terms into separate files and add them to tests.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    32
8c3cf9f4f5f2 Split Terms into separate files and add them to tests.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    33
local_setup {*
1277
6eacf60ce41d Permutation and FV_Alpha interface change.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1274
diff changeset
    34
  snd o define_fv_alpha (Datatype.the_info @{theory} "Term1.rtrm1")
1400
10eca65a8d03 alpha_eqvt for recursive term1.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1356
diff changeset
    35
  [[[], [], [(NONE, 0, 1)], [(SOME (@{term bv1}, true), 0, 2)]],
10eca65a8d03 alpha_eqvt for recursive term1.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1356
diff changeset
    36
  [[], [], []]] [(@{term bv1}, 1, [[], [0], [0, 1]])] *}
1270
8c3cf9f4f5f2 Split Terms into separate files and add them to tests.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    37
8c3cf9f4f5f2 Split Terms into separate files and add them to tests.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    38
notation
8c3cf9f4f5f2 Split Terms into separate files and add them to tests.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    39
  alpha_rtrm1 ("_ \<approx>1 _" [100, 100] 100) and
8c3cf9f4f5f2 Split Terms into separate files and add them to tests.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    40
  alpha_bp ("_ \<approx>1b _" [100, 100] 100)
1400
10eca65a8d03 alpha_eqvt for recursive term1.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1356
diff changeset
    41
thm alpha_rtrm1_alpha_bp_alpha_bv1.intros
1283
6a133abb7633 generated the "binding list" from the input; at the moment it is only printed out as tracing; does not yet include the "bind itself binders"
Christian Urban <urbanc@in.tum.de>
parents: 1278
diff changeset
    42
thm fv_rtrm1_fv_bp.simps[no_vars]
1270
8c3cf9f4f5f2 Split Terms into separate files and add them to tests.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    43
1400
10eca65a8d03 alpha_eqvt for recursive term1.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1356
diff changeset
    44
local_setup {* (fn ctxt => snd (Local_Theory.note ((@{binding alpha1_inj}, []), (build_alpha_inj @{thms alpha_rtrm1_alpha_bp_alpha_bv1.intros} @{thms rtrm1.distinct rtrm1.inject bp.distinct bp.inject} @{thms alpha_rtrm1.cases alpha_bp.cases alpha_bv1.cases} ctxt)) ctxt)) *}
1270
8c3cf9f4f5f2 Split Terms into separate files and add them to tests.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    45
thm alpha1_inj
8c3cf9f4f5f2 Split Terms into separate files and add them to tests.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    46
8c3cf9f4f5f2 Split Terms into separate files and add them to tests.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    47
local_setup {*
1425
112f998d2953 Fixes for term1 for new alpha. Still not able to show support equations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1400
diff changeset
    48
snd o (build_eqvts @{binding bv1_eqvt} [@{term bv1}] (build_eqvts_tac @{thm rtrm1_bp.inducts(2)} @{thms bv1.simps permute_rtrm1_permute_bp.simps} @{context}))
1270
8c3cf9f4f5f2 Split Terms into separate files and add them to tests.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    49
*}
8c3cf9f4f5f2 Split Terms into separate files and add them to tests.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    50
8c3cf9f4f5f2 Split Terms into separate files and add them to tests.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    51
local_setup {*
1425
112f998d2953 Fixes for term1 for new alpha. Still not able to show support equations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1400
diff changeset
    52
snd o build_eqvts @{binding fv_rtrm1_fv_bp_eqvt} [@{term fv_rtrm1}, @{term fv_bp}] (build_eqvts_tac @{thm rtrm1_bp.induct} @{thms fv_rtrm1_fv_bp.simps permute_rtrm1_permute_bp.simps} @{context})
1270
8c3cf9f4f5f2 Split Terms into separate files and add them to tests.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    53
*}
8c3cf9f4f5f2 Split Terms into separate files and add them to tests.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    54
1425
112f998d2953 Fixes for term1 for new alpha. Still not able to show support equations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1400
diff changeset
    55
lemma alpha1_eqvt: 
112f998d2953 Fixes for term1 for new alpha. Still not able to show support equations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1400
diff changeset
    56
  "(rtrm1 \<approx>1 rtrm1a \<longrightarrow> (p \<bullet> rtrm1) \<approx>1 (p \<bullet> rtrm1a)) \<and> 
112f998d2953 Fixes for term1 for new alpha. Still not able to show support equations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1400
diff changeset
    57
   (bp \<approx>1b bpa \<longrightarrow> (p \<bullet> bp) \<approx>1b (p \<bullet> bpa)) \<and>
112f998d2953 Fixes for term1 for new alpha. Still not able to show support equations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1400
diff changeset
    58
   (alpha_bv1 a b c \<longrightarrow> alpha_bv1 (p \<bullet> a) (p \<bullet> b) (p \<bullet> c))"
112f998d2953 Fixes for term1 for new alpha. Still not able to show support equations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1400
diff changeset
    59
by (tactic {* alpha_eqvt_tac @{thm alpha_rtrm1_alpha_bp_alpha_bv1.induct} @{thms alpha1_inj permute_rtrm1_permute_bp.simps} @{context} 1 *})
112f998d2953 Fixes for term1 for new alpha. Still not able to show support equations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1400
diff changeset
    60
1400
10eca65a8d03 alpha_eqvt for recursive term1.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1356
diff changeset
    61
(*
1300
22a084c9316b Fixed eqvt code.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1291
diff changeset
    62
local_setup {*
22a084c9316b Fixed eqvt code.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1291
diff changeset
    63
(fn ctxt => snd (Local_Theory.note ((@{binding alpha1_eqvt}, []),
1400
10eca65a8d03 alpha_eqvt for recursive term1.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1356
diff changeset
    64
build_alpha_eqvts [@{term alpha_rtrm1}, @{term alpha_bp}] [@{term "permute :: perm \<Rightarrow> rtrm1 \<Rightarrow> rtrm1"},@{term "permute :: perm \<Rightarrow> bp \<Rightarrow> bp"}] @{thms permute_rtrm1_permute_bp.simps alpha1_inj} @{thm alpha_rtrm1_alpha_bp_alpha_bv1.induct} ctxt) ctxt)) *}*)
1291
24889782da92 Trying to prove equivariance.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1283
diff changeset
    65
1270
8c3cf9f4f5f2 Split Terms into separate files and add them to tests.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    66
lemma alpha1_eqvt_proper[eqvt]:
8c3cf9f4f5f2 Split Terms into separate files and add them to tests.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    67
  "pi \<bullet> (t \<approx>1 s) = ((pi \<bullet> t) \<approx>1 (pi \<bullet> s))"
8c3cf9f4f5f2 Split Terms into separate files and add them to tests.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    68
  "pi \<bullet> (alpha_bp a b) = (alpha_bp (pi \<bullet> a) (pi \<bullet> b))"
8c3cf9f4f5f2 Split Terms into separate files and add them to tests.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    69
  apply (simp_all only: eqvts)
8c3cf9f4f5f2 Split Terms into separate files and add them to tests.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    70
  apply rule
8c3cf9f4f5f2 Split Terms into separate files and add them to tests.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    71
  apply (simp_all add: alpha1_eqvt)
8c3cf9f4f5f2 Split Terms into separate files and add them to tests.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    72
  apply (subst permute_minus_cancel(2)[symmetric,of "t" "pi"])
8c3cf9f4f5f2 Split Terms into separate files and add them to tests.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    73
  apply (subst permute_minus_cancel(2)[symmetric,of "s" "pi"])
8c3cf9f4f5f2 Split Terms into separate files and add them to tests.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    74
  apply (simp_all only: alpha1_eqvt)
8c3cf9f4f5f2 Split Terms into separate files and add them to tests.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    75
  apply rule
8c3cf9f4f5f2 Split Terms into separate files and add them to tests.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    76
  apply (simp_all add: alpha1_eqvt)
8c3cf9f4f5f2 Split Terms into separate files and add them to tests.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    77
  apply (subst permute_minus_cancel(2)[symmetric,of "a" "pi"])
8c3cf9f4f5f2 Split Terms into separate files and add them to tests.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    78
  apply (subst permute_minus_cancel(2)[symmetric,of "b" "pi"])
8c3cf9f4f5f2 Split Terms into separate files and add them to tests.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    79
  apply (simp_all only: alpha1_eqvt)
8c3cf9f4f5f2 Split Terms into separate files and add them to tests.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    80
done
1274
d867021d8ac1 Preparing the generalized lifting procedure
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1270
diff changeset
    81
thm eqvts_raw(1)
1270
8c3cf9f4f5f2 Split Terms into separate files and add them to tests.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    82
1425
112f998d2953 Fixes for term1 for new alpha. Still not able to show support equations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1400
diff changeset
    83
lemma alpha1_equivp:
112f998d2953 Fixes for term1 for new alpha. Still not able to show support equations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1400
diff changeset
    84
  "equivp alpha_rtrm1"
112f998d2953 Fixes for term1 for new alpha. Still not able to show support equations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1400
diff changeset
    85
  "equivp alpha_bp"
112f998d2953 Fixes for term1 for new alpha. Still not able to show support equations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1400
diff changeset
    86
sorry
112f998d2953 Fixes for term1 for new alpha. Still not able to show support equations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1400
diff changeset
    87
112f998d2953 Fixes for term1 for new alpha. Still not able to show support equations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1400
diff changeset
    88
(*
1270
8c3cf9f4f5f2 Split Terms into separate files and add them to tests.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    89
local_setup {* (fn ctxt => snd (Local_Theory.note ((@{binding alpha1_equivp}, []),
1400
10eca65a8d03 alpha_eqvt for recursive term1.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1356
diff changeset
    90
  (build_equivps [@{term alpha_rtrm1}, @{term alpha_bp}] @{thm rtrm1_bp.induct} @{thm alpha_rtrm1_alpha_bp_alpha_bv1.induct} @{thms rtrm1.inject bp.inject} @{thms alpha1_inj} @{thms rtrm1.distinct bp.distinct} @{thms alpha_rtrm1.cases alpha_bp.cases} @{thms alpha1_eqvt} ctxt)) ctxt)) *}
1270
8c3cf9f4f5f2 Split Terms into separate files and add them to tests.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    91
thm alpha1_equivp
1425
112f998d2953 Fixes for term1 for new alpha. Still not able to show support equations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1400
diff changeset
    92
*)
1270
8c3cf9f4f5f2 Split Terms into separate files and add them to tests.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    93
8c3cf9f4f5f2 Split Terms into separate files and add them to tests.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    94
local_setup  {* define_quotient_type [(([], @{binding trm1}, NoSyn), (@{typ rtrm1}, @{term alpha_rtrm1}))]
8c3cf9f4f5f2 Split Terms into separate files and add them to tests.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    95
  (rtac @{thm alpha1_equivp(1)} 1) *}
8c3cf9f4f5f2 Split Terms into separate files and add them to tests.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    96
8c3cf9f4f5f2 Split Terms into separate files and add them to tests.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    97
local_setup {*
8c3cf9f4f5f2 Split Terms into separate files and add them to tests.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    98
(fn ctxt => ctxt
8c3cf9f4f5f2 Split Terms into separate files and add them to tests.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    99
 |> snd o (Quotient_Def.quotient_lift_const ("Vr1", @{term rVr1}))
8c3cf9f4f5f2 Split Terms into separate files and add them to tests.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   100
 |> snd o (Quotient_Def.quotient_lift_const ("Ap1", @{term rAp1}))
8c3cf9f4f5f2 Split Terms into separate files and add them to tests.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   101
 |> snd o (Quotient_Def.quotient_lift_const ("Lm1", @{term rLm1}))
8c3cf9f4f5f2 Split Terms into separate files and add them to tests.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   102
 |> snd o (Quotient_Def.quotient_lift_const ("Lt1", @{term rLt1}))
8c3cf9f4f5f2 Split Terms into separate files and add them to tests.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   103
 |> snd o (Quotient_Def.quotient_lift_const ("fv_trm1", @{term fv_rtrm1})))
8c3cf9f4f5f2 Split Terms into separate files and add them to tests.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   104
*}
8c3cf9f4f5f2 Split Terms into separate files and add them to tests.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   105
print_theorems
8c3cf9f4f5f2 Split Terms into separate files and add them to tests.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   106
1278
8814494fe4da Change in signature of prove_const_rsp for general lifting.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1277
diff changeset
   107
local_setup {* snd o prove_const_rsp @{binding fv_rtrm1_rsp} [@{term fv_rtrm1}]
1425
112f998d2953 Fixes for term1 for new alpha. Still not able to show support equations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1400
diff changeset
   108
  (fn _ => Skip_Proof.cheat_tac @{theory}) *}
1278
8814494fe4da Change in signature of prove_const_rsp for general lifting.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1277
diff changeset
   109
local_setup {* snd o prove_const_rsp @{binding rVr1_rsp} [@{term rVr1}]
8814494fe4da Change in signature of prove_const_rsp for general lifting.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1277
diff changeset
   110
  (fn _ => constr_rsp_tac @{thms alpha1_inj} @{thms fv_rtrm1_rsp} @{thms alpha1_equivp} 1) *}
8814494fe4da Change in signature of prove_const_rsp for general lifting.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1277
diff changeset
   111
local_setup {* snd o prove_const_rsp @{binding rAp1_rsp} [@{term rAp1}]
1270
8c3cf9f4f5f2 Split Terms into separate files and add them to tests.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   112
  (fn _ => constr_rsp_tac @{thms alpha1_inj} @{thms fv_rtrm1_rsp} @{thms alpha1_equivp} 1) *}
1278
8814494fe4da Change in signature of prove_const_rsp for general lifting.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1277
diff changeset
   113
local_setup {* snd o prove_const_rsp @{binding rLm1_rsp} [@{term rLm1}]
1270
8c3cf9f4f5f2 Split Terms into separate files and add them to tests.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   114
  (fn _ => constr_rsp_tac @{thms alpha1_inj} @{thms fv_rtrm1_rsp} @{thms alpha1_equivp} 1) *}
1278
8814494fe4da Change in signature of prove_const_rsp for general lifting.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1277
diff changeset
   115
local_setup {* snd o prove_const_rsp @{binding rLt1_rsp} [@{term rLt1}]
1425
112f998d2953 Fixes for term1 for new alpha. Still not able to show support equations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1400
diff changeset
   116
  (fn _ => Skip_Proof.cheat_tac @{theory}) *}
1278
8814494fe4da Change in signature of prove_const_rsp for general lifting.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1277
diff changeset
   117
local_setup {* snd o prove_const_rsp @{binding permute_rtrm1_rsp} [@{term "permute :: perm \<Rightarrow> rtrm1 \<Rightarrow> rtrm1"}]
1270
8c3cf9f4f5f2 Split Terms into separate files and add them to tests.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   118
  (fn _ => asm_simp_tac (HOL_ss addsimps @{thms alpha1_eqvt}) 1) *}
8c3cf9f4f5f2 Split Terms into separate files and add them to tests.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   119
8c3cf9f4f5f2 Split Terms into separate files and add them to tests.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   120
lemmas trm1_bp_induct = rtrm1_bp.induct[quot_lifted]
8c3cf9f4f5f2 Split Terms into separate files and add them to tests.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   121
lemmas trm1_bp_inducts = rtrm1_bp.inducts[quot_lifted]
8c3cf9f4f5f2 Split Terms into separate files and add them to tests.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   122
8c3cf9f4f5f2 Split Terms into separate files and add them to tests.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   123
setup {* define_lifted_perms ["Term1.trm1"] [("permute_trm1", @{term "permute :: perm \<Rightarrow> rtrm1 \<Rightarrow> rtrm1"})]
8c3cf9f4f5f2 Split Terms into separate files and add them to tests.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   124
  @{thms permute_rtrm1_permute_bp_zero permute_rtrm1_permute_bp_append} *}
8c3cf9f4f5f2 Split Terms into separate files and add them to tests.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   125
8c3cf9f4f5f2 Split Terms into separate files and add them to tests.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   126
lemmas
8c3cf9f4f5f2 Split Terms into separate files and add them to tests.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   127
    permute_trm1 = permute_rtrm1_permute_bp.simps[quot_lifted]
8c3cf9f4f5f2 Split Terms into separate files and add them to tests.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   128
and fv_trm1 = fv_rtrm1_fv_bp.simps[quot_lifted]
8c3cf9f4f5f2 Split Terms into separate files and add them to tests.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   129
and fv_trm1_eqvt = fv_rtrm1_fv_bp_eqvt(1)[quot_lifted]
8c3cf9f4f5f2 Split Terms into separate files and add them to tests.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   130
and alpha1_INJ = alpha1_inj[unfolded alpha_gen, quot_lifted, folded alpha_gen]
8c3cf9f4f5f2 Split Terms into separate files and add them to tests.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   131
8c3cf9f4f5f2 Split Terms into separate files and add them to tests.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   132
lemma supports:
8c3cf9f4f5f2 Split Terms into separate files and add them to tests.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   133
  "(supp (atom x)) supports (Vr1 x)"
8c3cf9f4f5f2 Split Terms into separate files and add them to tests.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   134
  "(supp t \<union> supp s) supports (Ap1 t s)"
8c3cf9f4f5f2 Split Terms into separate files and add them to tests.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   135
  "(supp (atom x) \<union> supp t) supports (Lm1 x t)"
8c3cf9f4f5f2 Split Terms into separate files and add them to tests.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   136
  "(supp b \<union> supp t \<union> supp s) supports (Lt1 b t s)"
8c3cf9f4f5f2 Split Terms into separate files and add them to tests.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   137
  "{} supports BUnit"
8c3cf9f4f5f2 Split Terms into separate files and add them to tests.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   138
  "(supp (atom x)) supports (BVr x)"
8c3cf9f4f5f2 Split Terms into separate files and add them to tests.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   139
  "(supp a \<union> supp b) supports (BPr a b)"
8c3cf9f4f5f2 Split Terms into separate files and add them to tests.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   140
apply(simp_all add: supports_def fresh_def[symmetric] swap_fresh_fresh permute_trm1)
8c3cf9f4f5f2 Split Terms into separate files and add them to tests.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   141
apply(rule_tac [!] allI)+
8c3cf9f4f5f2 Split Terms into separate files and add them to tests.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   142
apply(rule_tac [!] impI)
8c3cf9f4f5f2 Split Terms into separate files and add them to tests.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   143
apply(tactic {* ALLGOALS (REPEAT o etac conjE) *})
8c3cf9f4f5f2 Split Terms into separate files and add them to tests.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   144
apply(simp_all add: fresh_atom)
8c3cf9f4f5f2 Split Terms into separate files and add them to tests.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   145
done
8c3cf9f4f5f2 Split Terms into separate files and add them to tests.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   146
8c3cf9f4f5f2 Split Terms into separate files and add them to tests.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   147
lemma rtrm1_bp_fs:
8c3cf9f4f5f2 Split Terms into separate files and add them to tests.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   148
  "finite (supp (x :: trm1))"
8c3cf9f4f5f2 Split Terms into separate files and add them to tests.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   149
  "finite (supp (b :: bp))"
8c3cf9f4f5f2 Split Terms into separate files and add them to tests.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   150
  apply (induct x and b rule: trm1_bp_inducts)
8c3cf9f4f5f2 Split Terms into separate files and add them to tests.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   151
  apply(tactic {* ALLGOALS (rtac @{thm supports_finite} THEN' resolve_tac @{thms supports}) *})
8c3cf9f4f5f2 Split Terms into separate files and add them to tests.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   152
  apply(simp_all add: supp_atom)
8c3cf9f4f5f2 Split Terms into separate files and add them to tests.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   153
  done
8c3cf9f4f5f2 Split Terms into separate files and add them to tests.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   154
1343
693df83172f0 Trying to fix the proofs for the single existential... So far failed.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1300
diff changeset
   155
instance trm1 and bp :: fs
1270
8c3cf9f4f5f2 Split Terms into separate files and add them to tests.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   156
apply default
1343
693df83172f0 Trying to fix the proofs for the single existential... So far failed.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1300
diff changeset
   157
apply (rule rtrm1_bp_fs)+
1270
8c3cf9f4f5f2 Split Terms into separate files and add them to tests.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   158
done
1425
112f998d2953 Fixes for term1 for new alpha. Still not able to show support equations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1400
diff changeset
   159
1345
8d2667ebe26c Not much progress about the single existential let case.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1343
diff changeset
   160
lemma fv_eq_bv_pre: "fv_bp bp = bv1 bp"
1270
8c3cf9f4f5f2 Split Terms into separate files and add them to tests.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   161
apply(induct bp rule: trm1_bp_inducts(2))
8c3cf9f4f5f2 Split Terms into separate files and add them to tests.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   162
apply(simp_all)
8c3cf9f4f5f2 Split Terms into separate files and add them to tests.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   163
done
8c3cf9f4f5f2 Split Terms into separate files and add them to tests.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   164
1345
8d2667ebe26c Not much progress about the single existential let case.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1343
diff changeset
   165
lemma fv_eq_bv: "fv_bp = bv1"
8d2667ebe26c Not much progress about the single existential let case.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1343
diff changeset
   166
apply(rule ext)
8d2667ebe26c Not much progress about the single existential let case.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1343
diff changeset
   167
apply(rule fv_eq_bv_pre)
8d2667ebe26c Not much progress about the single existential let case.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1343
diff changeset
   168
done
8d2667ebe26c Not much progress about the single existential let case.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1343
diff changeset
   169
1270
8c3cf9f4f5f2 Split Terms into separate files and add them to tests.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   170
lemma helper2: "{b. \<forall>pi. pi \<bullet> (a \<rightleftharpoons> b) \<bullet> bp \<noteq> bp} = {}"
8c3cf9f4f5f2 Split Terms into separate files and add them to tests.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   171
apply auto
8c3cf9f4f5f2 Split Terms into separate files and add them to tests.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   172
apply (rule_tac x="(x \<rightleftharpoons> a)" in exI)
8c3cf9f4f5f2 Split Terms into separate files and add them to tests.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   173
apply auto
8c3cf9f4f5f2 Split Terms into separate files and add them to tests.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   174
done
8c3cf9f4f5f2 Split Terms into separate files and add them to tests.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   175
8c3cf9f4f5f2 Split Terms into separate files and add them to tests.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   176
lemma alpha_bp_eq_eq: "alpha_bp a b = (a = b)"
8c3cf9f4f5f2 Split Terms into separate files and add them to tests.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   177
apply rule
1425
112f998d2953 Fixes for term1 for new alpha. Still not able to show support equations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1400
diff changeset
   178
apply (induct a b rule: alpha_rtrm1_alpha_bp_alpha_bv1.inducts(2))
1270
8c3cf9f4f5f2 Split Terms into separate files and add them to tests.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   179
apply (simp_all add: equivp_reflp[OF alpha1_equivp(2)])
8c3cf9f4f5f2 Split Terms into separate files and add them to tests.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   180
done
8c3cf9f4f5f2 Split Terms into separate files and add them to tests.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   181
8c3cf9f4f5f2 Split Terms into separate files and add them to tests.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   182
lemma alpha_bp_eq: "alpha_bp = (op =)"
8c3cf9f4f5f2 Split Terms into separate files and add them to tests.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   183
apply (rule ext)+
8c3cf9f4f5f2 Split Terms into separate files and add them to tests.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   184
apply (rule alpha_bp_eq_eq)
8c3cf9f4f5f2 Split Terms into separate files and add them to tests.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   185
done
8c3cf9f4f5f2 Split Terms into separate files and add them to tests.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   186
1345
8d2667ebe26c Not much progress about the single existential let case.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1343
diff changeset
   187
lemma ex_out: 
8d2667ebe26c Not much progress about the single existential let case.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1343
diff changeset
   188
  "(\<exists>x. Z x \<and> Q) = (Q \<and> (\<exists>x. Z x))"
8d2667ebe26c Not much progress about the single existential let case.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1343
diff changeset
   189
  "(\<exists>x. Q \<and> Z x) = (Q \<and> (\<exists>x. Z x))"
8d2667ebe26c Not much progress about the single existential let case.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1343
diff changeset
   190
  "(\<exists>x. P x \<and> Q \<and> Z x) = (Q \<and> (\<exists>x. P x \<and> Z x))"
8d2667ebe26c Not much progress about the single existential let case.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1343
diff changeset
   191
  "(\<exists>x. Q \<and> P x \<and> Z x) = (Q \<and> (\<exists>x. P x \<and> Z x))"
8d2667ebe26c Not much progress about the single existential let case.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1343
diff changeset
   192
  "(\<exists>x. Q \<and> P x \<and> Z x \<and> W x) = (Q \<and> (\<exists>x. P x \<and> Z x \<and> W x))"
8d2667ebe26c Not much progress about the single existential let case.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1343
diff changeset
   193
apply (blast)+
8d2667ebe26c Not much progress about the single existential let case.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1343
diff changeset
   194
done
8d2667ebe26c Not much progress about the single existential let case.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1343
diff changeset
   195
8d2667ebe26c Not much progress about the single existential let case.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1343
diff changeset
   196
lemma "(Abs bs (x, x') = Abs cs (y, y')) = (\<exists>p. (bs, x) \<approx>gen op = supp p (cs, y) \<and> (bs, x') \<approx>gen op = supp p (cs, y'))"
8d2667ebe26c Not much progress about the single existential let case.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1343
diff changeset
   197
thm Abs_eq_iff
8d2667ebe26c Not much progress about the single existential let case.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1343
diff changeset
   198
apply (simp add: Abs_eq_iff)
8d2667ebe26c Not much progress about the single existential let case.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1343
diff changeset
   199
apply (rule arg_cong[of _ _ "Ex"])
8d2667ebe26c Not much progress about the single existential let case.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1343
diff changeset
   200
apply (rule ext)
8d2667ebe26c Not much progress about the single existential let case.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1343
diff changeset
   201
apply (simp only: alpha_gen)
8d2667ebe26c Not much progress about the single existential let case.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1343
diff changeset
   202
apply (simp only: supp_Pair eqvts)
8d2667ebe26c Not much progress about the single existential let case.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1343
diff changeset
   203
apply rule
8d2667ebe26c Not much progress about the single existential let case.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1343
diff changeset
   204
apply (erule conjE)+
8d2667ebe26c Not much progress about the single existential let case.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1343
diff changeset
   205
oops
8d2667ebe26c Not much progress about the single existential let case.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1343
diff changeset
   206
8d2667ebe26c Not much progress about the single existential let case.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1343
diff changeset
   207
lemma "(f (p \<bullet> bp), p \<bullet> bp) \<approx>gen op = f pi (f bp, bp) = False"
8d2667ebe26c Not much progress about the single existential let case.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1343
diff changeset
   208
apply (simp add: alpha_gen fresh_star_def)
8d2667ebe26c Not much progress about the single existential let case.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1343
diff changeset
   209
oops
8d2667ebe26c Not much progress about the single existential let case.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1343
diff changeset
   210
8d2667ebe26c Not much progress about the single existential let case.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1343
diff changeset
   211
(* TODO: permute_ABS should be in eqvt? *)
8d2667ebe26c Not much progress about the single existential let case.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1343
diff changeset
   212
8d2667ebe26c Not much progress about the single existential let case.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1343
diff changeset
   213
lemma Collect_neg_conj: "{x. \<not>(P x \<and> Q x)} = {x. \<not>(P x)} \<union> {x. \<not>(Q x)}"
8d2667ebe26c Not much progress about the single existential let case.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1343
diff changeset
   214
by (simp add: Collect_imp_eq Collect_neg_eq[symmetric])
8d2667ebe26c Not much progress about the single existential let case.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1343
diff changeset
   215
1349
6204137160d8 Still unable to show supp=fv for let with one existential.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1345
diff changeset
   216
lemma "
6204137160d8 Still unable to show supp=fv for let with one existential.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1345
diff changeset
   217
{a\<Colon>atom. infinite ({b\<Colon>atom. \<not> (\<exists>pi\<Colon>perm. P pi a b \<and> Q pi a b)})} =
6204137160d8 Still unable to show supp=fv for let with one existential.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1345
diff changeset
   218
{a\<Colon>atom. infinite {b\<Colon>atom. \<not> (\<exists>p\<Colon>perm. P p a b)}} \<union>
6204137160d8 Still unable to show supp=fv for let with one existential.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1345
diff changeset
   219
{a\<Colon>atom. infinite {b\<Colon>atom. \<not> (\<exists>p\<Colon>perm. Q p a b)}}"
6204137160d8 Still unable to show supp=fv for let with one existential.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1345
diff changeset
   220
oops
6204137160d8 Still unable to show supp=fv for let with one existential.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1345
diff changeset
   221
6204137160d8 Still unable to show supp=fv for let with one existential.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1345
diff changeset
   222
lemma inf_or: "(infinite x \<or> infinite y) = infinite (x \<union> y)"
6204137160d8 Still unable to show supp=fv for let with one existential.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1345
diff changeset
   223
by (simp add: finite_Un)
6204137160d8 Still unable to show supp=fv for let with one existential.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1345
diff changeset
   224
6204137160d8 Still unable to show supp=fv for let with one existential.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1345
diff changeset
   225
6204137160d8 Still unable to show supp=fv for let with one existential.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1345
diff changeset
   226
lemma supp_fv_let:
6204137160d8 Still unable to show supp=fv for let with one existential.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1345
diff changeset
   227
  assumes sa : "fv_bp bp = supp bp"
1425
112f998d2953 Fixes for term1 for new alpha. Still not able to show support equations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1400
diff changeset
   228
  shows "\<lbrakk>fv_trm1 trm11 = supp trm11; fv_trm1 trm12 = supp trm12\<rbrakk>
112f998d2953 Fixes for term1 for new alpha. Still not able to show support equations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1400
diff changeset
   229
       \<Longrightarrow> supp (Lt1 bp trm11 trm12) = fv_trm1 (Lt1 bp trm11 trm12)"
1349
6204137160d8 Still unable to show supp=fv for let with one existential.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1345
diff changeset
   230
apply(simp only: fv_trm1 fv_eq_bv sa[simplified fv_eq_bv])
6204137160d8 Still unable to show supp=fv for let with one existential.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1345
diff changeset
   231
apply(fold supp_Abs)
6204137160d8 Still unable to show supp=fv for let with one existential.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1345
diff changeset
   232
apply(simp only: fv_trm1 fv_eq_bv sa[simplified fv_eq_bv,symmetric])
6204137160d8 Still unable to show supp=fv for let with one existential.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1345
diff changeset
   233
apply(simp (no_asm) only: supp_def permute_set_eq permute_trm1 alpha1_INJ)
6204137160d8 Still unable to show supp=fv for let with one existential.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1345
diff changeset
   234
apply(simp only: ex_out Collect_neg_conj permute_ABS Abs_eq_iff)
1425
112f998d2953 Fixes for term1 for new alpha. Still not able to show support equations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1400
diff changeset
   235
(*apply(simp only: alpha_bp_eq fv_eq_bv)*)
1349
6204137160d8 Still unable to show supp=fv for let with one existential.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1345
diff changeset
   236
apply(simp only: alpha_gen fv_eq_bv supp_Pair)
6204137160d8 Still unable to show supp=fv for let with one existential.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1345
diff changeset
   237
apply(simp only: supp_eqvt[symmetric] fv_trm1_eqvt[symmetric] bv1_eqvt fv_eq_bv sa[simplified fv_eq_bv,symmetric])
6204137160d8 Still unable to show supp=fv for let with one existential.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1345
diff changeset
   238
apply(simp only: Un_left_commute)
6204137160d8 Still unable to show supp=fv for let with one existential.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1345
diff changeset
   239
apply simp
6204137160d8 Still unable to show supp=fv for let with one existential.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1345
diff changeset
   240
apply(simp add: fresh_star_def) apply(fold fresh_star_def)
6204137160d8 Still unable to show supp=fv for let with one existential.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1345
diff changeset
   241
apply(simp add: Collect_imp_eq Collect_neg_eq[symmetric])
6204137160d8 Still unable to show supp=fv for let with one existential.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1345
diff changeset
   242
apply(tactic {* Cong_Tac.cong_tac @{thm cong} 1 *}) apply(rule refl)
6204137160d8 Still unable to show supp=fv for let with one existential.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1345
diff changeset
   243
apply(simp only: Un_assoc[symmetric])
6204137160d8 Still unable to show supp=fv for let with one existential.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1345
diff changeset
   244
apply(simp only: Un_commute)
6204137160d8 Still unable to show supp=fv for let with one existential.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1345
diff changeset
   245
apply(simp only: Un_left_commute)
6204137160d8 Still unable to show supp=fv for let with one existential.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1345
diff changeset
   246
apply(simp only: Un_assoc[symmetric])
6204137160d8 Still unable to show supp=fv for let with one existential.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1345
diff changeset
   247
apply(simp only: Un_commute)
6204137160d8 Still unable to show supp=fv for let with one existential.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1345
diff changeset
   248
apply(tactic {* Cong_Tac.cong_tac @{thm cong} 1 *}) apply(rule refl)
6204137160d8 Still unable to show supp=fv for let with one existential.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1345
diff changeset
   249
apply(simp only: Collect_disj_eq[symmetric] inf_or)
1356
094811120a68 Undo effects of simp.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1349
diff changeset
   250
apply(simp only: HOL.imp_disj2[symmetric] HOL.not_ex[symmetric] HOL.disj_not1[symmetric] HOL.de_Morgan_conj[symmetric])
1349
6204137160d8 Still unable to show supp=fv for let with one existential.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1345
diff changeset
   251
sorry
6204137160d8 Still unable to show supp=fv for let with one existential.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1345
diff changeset
   252
1270
8c3cf9f4f5f2 Split Terms into separate files and add them to tests.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   253
lemma supp_fv:
8c3cf9f4f5f2 Split Terms into separate files and add them to tests.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   254
  "supp t = fv_trm1 t"
8c3cf9f4f5f2 Split Terms into separate files and add them to tests.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   255
  "supp b = fv_bp b"
8c3cf9f4f5f2 Split Terms into separate files and add them to tests.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   256
apply(induct t and b rule: trm1_bp_inducts)
8c3cf9f4f5f2 Split Terms into separate files and add them to tests.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   257
apply(simp_all)
8c3cf9f4f5f2 Split Terms into separate files and add them to tests.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   258
apply(simp add: supp_def permute_trm1 alpha1_INJ fv_trm1)
8c3cf9f4f5f2 Split Terms into separate files and add them to tests.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   259
apply(simp only: supp_at_base[simplified supp_def])
8c3cf9f4f5f2 Split Terms into separate files and add them to tests.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   260
apply(simp add: supp_def permute_trm1 alpha1_INJ fv_trm1)
1343
693df83172f0 Trying to fix the proofs for the single existential... So far failed.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1300
diff changeset
   261
apply(simp add: Collect_imp_eq Collect_neg_eq Un_commute)
1270
8c3cf9f4f5f2 Split Terms into separate files and add them to tests.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   262
apply(subgoal_tac "supp (Lm1 name rtrm1) = supp (Abs {atom name} rtrm1)")
8c3cf9f4f5f2 Split Terms into separate files and add them to tests.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   263
apply(simp add: supp_Abs fv_trm1)
8c3cf9f4f5f2 Split Terms into separate files and add them to tests.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   264
apply(simp (no_asm) add: supp_def permute_set_eq atom_eqvt permute_trm1)
8c3cf9f4f5f2 Split Terms into separate files and add them to tests.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   265
apply(simp add: alpha1_INJ)
8c3cf9f4f5f2 Split Terms into separate files and add them to tests.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   266
apply(simp add: Abs_eq_iff)
8c3cf9f4f5f2 Split Terms into separate files and add them to tests.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   267
apply(simp add: alpha_gen.simps)
8c3cf9f4f5f2 Split Terms into separate files and add them to tests.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   268
apply(simp add: supp_eqvt[symmetric] fv_trm1_eqvt[symmetric])
1343
693df83172f0 Trying to fix the proofs for the single existential... So far failed.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1300
diff changeset
   269
defer
1270
8c3cf9f4f5f2 Split Terms into separate files and add them to tests.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   270
apply(simp (no_asm) add: supp_def permute_set_eq atom_eqvt)
8c3cf9f4f5f2 Split Terms into separate files and add them to tests.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   271
apply(simp (no_asm) add: supp_def eqvts)
8c3cf9f4f5f2 Split Terms into separate files and add them to tests.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   272
apply(fold supp_def)
8c3cf9f4f5f2 Split Terms into separate files and add them to tests.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   273
apply(simp add: supp_at_base)
8c3cf9f4f5f2 Split Terms into separate files and add them to tests.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   274
apply(simp (no_asm) add: supp_def Collect_imp_eq Collect_neg_eq)
8c3cf9f4f5f2 Split Terms into separate files and add them to tests.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   275
apply(simp add: Collect_imp_eq[symmetric] Collect_neg_eq[symmetric] supp_def[symmetric])
1349
6204137160d8 Still unable to show supp=fv for let with one existential.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1345
diff changeset
   276
(*apply(rule supp_fv_let) apply(simp_all)*)
1345
8d2667ebe26c Not much progress about the single existential let case.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1343
diff changeset
   277
apply(subgoal_tac "supp (Lt1 bp rtrm11 rtrm12) = supp (Abs (bv1 bp) (rtrm12)) \<union> supp(rtrm11)")
8d2667ebe26c Not much progress about the single existential let case.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1343
diff changeset
   278
(*apply(subgoal_tac "supp (Lt1 bp rtrm11 rtrm12) = supp (Abs (bv1 bp) (bp, rtrm12)) \<union> supp(rtrm11)")*)
1343
693df83172f0 Trying to fix the proofs for the single existential... So far failed.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1300
diff changeset
   279
apply(simp add: supp_Abs fv_trm1 supp_Pair Un_Diff Un_assoc fv_eq_bv)
693df83172f0 Trying to fix the proofs for the single existential... So far failed.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1300
diff changeset
   280
apply(blast) (* Un_commute in a good place *)
1345
8d2667ebe26c Not much progress about the single existential let case.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1343
diff changeset
   281
apply(simp (no_asm) only: supp_def permute_set_eq atom_eqvt permute_trm1)
8d2667ebe26c Not much progress about the single existential let case.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1343
diff changeset
   282
apply(simp only: alpha1_INJ permute_ABS permute_prod.simps Abs_eq_iff)
8d2667ebe26c Not much progress about the single existential let case.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1343
diff changeset
   283
apply(simp only: ex_out)
1343
693df83172f0 Trying to fix the proofs for the single existential... So far failed.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1300
diff changeset
   284
apply(simp only: Un_commute)
1345
8d2667ebe26c Not much progress about the single existential let case.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1343
diff changeset
   285
apply(simp only: alpha_bp_eq fv_eq_bv)
8d2667ebe26c Not much progress about the single existential let case.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1343
diff changeset
   286
apply(simp only: alpha_gen fv_eq_bv supp_Pair)
8d2667ebe26c Not much progress about the single existential let case.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1343
diff changeset
   287
apply(simp only: supp_eqvt[symmetric] fv_trm1_eqvt[symmetric] bv1_eqvt fv_eq_bv)
8d2667ebe26c Not much progress about the single existential let case.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1343
diff changeset
   288
apply(simp only: ex_out)
8d2667ebe26c Not much progress about the single existential let case.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1343
diff changeset
   289
apply(simp only: Collect_neg_conj finite_Un Diff_cancel)
8d2667ebe26c Not much progress about the single existential let case.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1343
diff changeset
   290
apply(simp)
1343
693df83172f0 Trying to fix the proofs for the single existential... So far failed.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1300
diff changeset
   291
apply(simp add: Collect_imp_eq)
1345
8d2667ebe26c Not much progress about the single existential let case.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1343
diff changeset
   292
apply(simp add: Collect_neg_eq[symmetric] fresh_star_def)
8d2667ebe26c Not much progress about the single existential let case.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1343
diff changeset
   293
apply(fold supp_def)
1343
693df83172f0 Trying to fix the proofs for the single existential... So far failed.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1300
diff changeset
   294
sorry
1270
8c3cf9f4f5f2 Split Terms into separate files and add them to tests.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   295
8c3cf9f4f5f2 Split Terms into separate files and add them to tests.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   296
lemma trm1_supp:
8c3cf9f4f5f2 Split Terms into separate files and add them to tests.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   297
  "supp (Vr1 x) = {atom x}"
8c3cf9f4f5f2 Split Terms into separate files and add them to tests.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   298
  "supp (Ap1 t1 t2) = supp t1 \<union> supp t2"
8c3cf9f4f5f2 Split Terms into separate files and add them to tests.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   299
  "supp (Lm1 x t) = (supp t) - {atom x}"
8c3cf9f4f5f2 Split Terms into separate files and add them to tests.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   300
  "supp (Lt1 b t s) = supp t \<union> (supp s - bv1 b)"
8c3cf9f4f5f2 Split Terms into separate files and add them to tests.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   301
by (simp_all add: supp_fv fv_trm1 fv_eq_bv)
8c3cf9f4f5f2 Split Terms into separate files and add them to tests.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   302
8c3cf9f4f5f2 Split Terms into separate files and add them to tests.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   303
lemma trm1_induct_strong:
8c3cf9f4f5f2 Split Terms into separate files and add them to tests.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   304
  assumes "\<And>name b. P b (Vr1 name)"
8c3cf9f4f5f2 Split Terms into separate files and add them to tests.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   305
  and     "\<And>rtrm11 rtrm12 b. \<lbrakk>\<And>c. P c rtrm11; \<And>c. P c rtrm12\<rbrakk> \<Longrightarrow> P b (Ap1 rtrm11 rtrm12)"
8c3cf9f4f5f2 Split Terms into separate files and add them to tests.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   306
  and     "\<And>name rtrm1 b. \<lbrakk>\<And>c. P c rtrm1; (atom name) \<sharp> b\<rbrakk> \<Longrightarrow> P b (Lm1 name rtrm1)"
8c3cf9f4f5f2 Split Terms into separate files and add them to tests.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   307
  and     "\<And>bp rtrm11 rtrm12 b. \<lbrakk>\<And>c. P c rtrm11; \<And>c. P c rtrm12; bv1 bp \<sharp>* b\<rbrakk> \<Longrightarrow> P b (Lt1 bp rtrm11 rtrm12)"
8c3cf9f4f5f2 Split Terms into separate files and add them to tests.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   308
  shows   "P a rtrma"
8c3cf9f4f5f2 Split Terms into separate files and add them to tests.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   309
sorry
8c3cf9f4f5f2 Split Terms into separate files and add them to tests.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   310
8c3cf9f4f5f2 Split Terms into separate files and add them to tests.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   311
end