src/Pure/ex/Guess_Examples.thy
author wenzelm
Fri, 04 Mar 2022 23:22:39 +0100
changeset 75215 1129e82dc1ec
parent 74365 b49bd5d9041f
permissions -rw-r--r--
more robust;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
74365
b49bd5d9041f improper proof command 'guess' moved to separate theory "Pure-ex.Guess";
wenzelm
parents: 61343
diff changeset
     1
(*  Title:      Pure/ex/Guess_Examples.thy
20005
3fd6d57b16de Proof by guessing.
wenzelm
parents:
diff changeset
     2
    Author:     Makarius
3fd6d57b16de Proof by guessing.
wenzelm
parents:
diff changeset
     3
*)
3fd6d57b16de Proof by guessing.
wenzelm
parents:
diff changeset
     4
74365
b49bd5d9041f improper proof command 'guess' moved to separate theory "Pure-ex.Guess";
wenzelm
parents: 61343
diff changeset
     5
section \<open>Proof by wild guessing\<close>
20005
3fd6d57b16de Proof by guessing.
wenzelm
parents:
diff changeset
     6
74365
b49bd5d9041f improper proof command 'guess' moved to separate theory "Pure-ex.Guess";
wenzelm
parents: 61343
diff changeset
     7
theory Guess_Examples
b49bd5d9041f improper proof command 'guess' moved to separate theory "Pure-ex.Guess";
wenzelm
parents: 61343
diff changeset
     8
imports "Pure-Examples.Higher_Order_Logic" Guess
20005
3fd6d57b16de Proof by guessing.
wenzelm
parents:
diff changeset
     9
begin
3fd6d57b16de Proof by guessing.
wenzelm
parents:
diff changeset
    10
74365
b49bd5d9041f improper proof command 'guess' moved to separate theory "Pure-ex.Guess";
wenzelm
parents: 61343
diff changeset
    11
typedecl t
b49bd5d9041f improper proof command 'guess' moved to separate theory "Pure-ex.Guess";
wenzelm
parents: 61343
diff changeset
    12
instance t :: type ..
b49bd5d9041f improper proof command 'guess' moved to separate theory "Pure-ex.Guess";
wenzelm
parents: 61343
diff changeset
    13
59031
4c3bb56b8ce7 misc tuning and modernization;
wenzelm
parents: 58889
diff changeset
    14
notepad
4c3bb56b8ce7 misc tuning and modernization;
wenzelm
parents: 58889
diff changeset
    15
begin
74365
b49bd5d9041f improper proof command 'guess' moved to separate theory "Pure-ex.Guess";
wenzelm
parents: 61343
diff changeset
    16
  have 1: "\<exists>x :: 'a. x = x" by (intro exI refl)
20005
3fd6d57b16de Proof by guessing.
wenzelm
parents:
diff changeset
    17
3fd6d57b16de Proof by guessing.
wenzelm
parents:
diff changeset
    18
  from 1 guess ..
3fd6d57b16de Proof by guessing.
wenzelm
parents:
diff changeset
    19
  from 1 guess x ..
3fd6d57b16de Proof by guessing.
wenzelm
parents:
diff changeset
    20
  from 1 guess x :: 'a ..
74365
b49bd5d9041f improper proof command 'guess' moved to separate theory "Pure-ex.Guess";
wenzelm
parents: 61343
diff changeset
    21
  from 1 guess x :: t ..
20005
3fd6d57b16de Proof by guessing.
wenzelm
parents:
diff changeset
    22
74365
b49bd5d9041f improper proof command 'guess' moved to separate theory "Pure-ex.Guess";
wenzelm
parents: 61343
diff changeset
    23
  have 2: "\<exists>(x::'c) (y::'d). x = x \<and> y = y" by (intro exI conjI refl)
20005
3fd6d57b16de Proof by guessing.
wenzelm
parents:
diff changeset
    24
  from 2 guess apply - apply (erule exE conjE)+ done
3fd6d57b16de Proof by guessing.
wenzelm
parents:
diff changeset
    25
  from 2 guess x apply - apply (erule exE conjE)+ done
3fd6d57b16de Proof by guessing.
wenzelm
parents:
diff changeset
    26
  from 2 guess x y apply - apply (erule exE conjE)+ done
3fd6d57b16de Proof by guessing.
wenzelm
parents:
diff changeset
    27
  from 2 guess x :: 'a and y :: 'b apply - apply (erule exE conjE)+ done
74365
b49bd5d9041f improper proof command 'guess' moved to separate theory "Pure-ex.Guess";
wenzelm
parents: 61343
diff changeset
    28
  from 2 guess x y :: t apply - apply (erule exE conjE)+ done
59031
4c3bb56b8ce7 misc tuning and modernization;
wenzelm
parents: 58889
diff changeset
    29
end
20005
3fd6d57b16de Proof by guessing.
wenzelm
parents:
diff changeset
    30
3fd6d57b16de Proof by guessing.
wenzelm
parents:
diff changeset
    31
end