Nominal/Ex/FiniteType.thy
author Christian Urban <urbanc@in.tum.de>
Thu, 04 Oct 2012 11:10:23 +0100
changeset 3200 995d47b09ab4
parent 3109 d79e936e30ea
child 3235 5ebd327ffb96
permissions -rw-r--r--
removed fork_mono flag
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
3109
d79e936e30ea commented out parts of TypeScheme1 in order to run all tests
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     1
theory FiniteType
d79e936e30ea commented out parts of TypeScheme1 in order to run all tests
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     2
imports "../Nominal2"
d79e936e30ea commented out parts of TypeScheme1 in order to run all tests
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     3
begin
d79e936e30ea commented out parts of TypeScheme1 in order to run all tests
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     4
d79e936e30ea commented out parts of TypeScheme1 in order to run all tests
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     5
typedef ('a::pt, 'b::fs) ffun = "{f::'a => 'b. finite (supp f)}"
d79e936e30ea commented out parts of TypeScheme1 in order to run all tests
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     6
apply(rule_tac x="\<lambda>_. undefined::'b::fs" in exI)
d79e936e30ea commented out parts of TypeScheme1 in order to run all tests
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     7
apply(auto)
d79e936e30ea commented out parts of TypeScheme1 in order to run all tests
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     8
apply(rule_tac S="supp (undefined::'b::fs)" in supports_finite)
d79e936e30ea commented out parts of TypeScheme1 in order to run all tests
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     9
apply(simp add: supports_def)
d79e936e30ea commented out parts of TypeScheme1 in order to run all tests
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    10
apply(perm_simp)
d79e936e30ea commented out parts of TypeScheme1 in order to run all tests
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    11
apply(simp add: fresh_def[symmetric])
d79e936e30ea commented out parts of TypeScheme1 in order to run all tests
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    12
apply(simp add: swap_fresh_fresh)
d79e936e30ea commented out parts of TypeScheme1 in order to run all tests
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    13
apply(simp add: finite_supp)
d79e936e30ea commented out parts of TypeScheme1 in order to run all tests
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    14
done
d79e936e30ea commented out parts of TypeScheme1 in order to run all tests
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    15
d79e936e30ea commented out parts of TypeScheme1 in order to run all tests
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    16
d79e936e30ea commented out parts of TypeScheme1 in order to run all tests
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    17
d79e936e30ea commented out parts of TypeScheme1 in order to run all tests
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    18
d79e936e30ea commented out parts of TypeScheme1 in order to run all tests
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    19
end